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.
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) Dintros; 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, included (interior D) Dforall 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, open_set D -> included D (interior D)forall D : R -> Prop, included D (adherence D)forall D : R -> Prop, included D (adherence D)D:R -> Propx:RH:D xV:R -> PropH0:neighbourhood V xV xD:R -> Propx:RH:D xV:R -> PropH0:neighbourhood V xD xapply H. Qed.D:R -> Propx:RH:D xV:R -> PropH0:neighbourhood V xD xforall D1 D2 D3 : R -> Prop, included D1 D2 -> included D2 D3 -> included D1 D3unfold included; intros; apply H0; apply H; apply H1. Qed.forall D1 D2 D3 : R -> Prop, included D1 D2 -> included D2 D3 -> included D1 D3forall D : R -> Prop, open_set (interior D)forall D : R -> Prop, open_set (interior D)D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dexists delta : posreal, included (disc x delta) (fun x1 : R => exists delta0 : posreal, included (disc x1 delta0) D)D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1exists delta : posreal, forall x2 : R, disc x1 delta x2 -> D x2D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):Rexists delta : posreal, forall x2 : R, disc x1 delta x2 -> D x2D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < del -> exists delta : posreal, forall x2 : R, disc x1 delta x2 -> D x2D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2D x2D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0) -> D x2D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0)D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2H4:included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0)H5:included (disc x1 {| pos := del; cond_pos := H2 |}) DD x2D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0)D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0)D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2x3:RH4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}Rabs (x3 - x) < x0D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2x3:RH4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}Rabs (x3 - x) <= Rabs (x3 - x1) + Rabs (x1 - x)D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2x3:RH4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}Rabs (x3 - x1) + Rabs (x1 - x) < x0D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2x3:RH4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}Rabs (x3 - x1) + Rabs (x1 - x) < x0D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2x3:RH4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}Rabs (x3 - x1) + Rabs (x1 - x) < del + Rabs (x1 - x)D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2x3:RH4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}del + Rabs (x1 - x) = x0D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RH2:0 < delx2:RH3:disc x1 {| pos := del; cond_pos := H2 |} x2x3:RH4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}del + Rabs (x1 - x) = x0D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delD:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):R0 < delunfold disc in H1; rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply H1. Qed.D:R -> Propx:RH:exists delta : posreal, included (disc x delta) Dx0:posrealH0:included (disc x x0) Dx1:RH1:disc x x0 x1del:=x0 - Rabs (x - x1):RRabs (x - x1) < x0forall 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, ~ (exists y : R, intersection_domain D (complementary D) y)forall D : R -> Prop, closed_set D -> included (adherence D) Dforall D : R -> Prop, closed_set D -> included (adherence D) DD:R -> PropH:forall x0 : R, ~ D x0 -> neighbourhood (fun c : R => ~ D c) x0x:RH0:point_adherent D xH1:D x \/ ~ D xH2:D xD xD:R -> PropH:forall x0 : R, ~ D x0 -> neighbourhood (fun c : R => ~ D c) x0x:RH0:point_adherent D xH1:D x \/ ~ D xH2:~ D xD xassert (H3 := H _ H2); assert (H4 := H0 _ H3); elim H4; intros; unfold intersection_domain in H5; elim H5; intros; elim H6; assumption. Qed.D:R -> PropH:forall x0 : R, ~ D x0 -> neighbourhood (fun c : R => ~ D c) x0x:RH0:point_adherent D xH1:D x \/ ~ D xH2:~ D xD xforall D : R -> Prop, closed_set (adherence D)forall D : R -> Prop, closed_set (adherence D)D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yFalseD:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(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) -> FalseD:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D yexists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D yexists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y0 < x0 - Rabs (x - x1) -> exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y0 < x0 - Rabs (x - x1)D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:forall x3 : R, disc x x0 x3 -> V0 x3x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D yH9:0 < x0 - Rabs (x - x1)del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posrealx2:RH10:disc x1 del x2Rabs (x2 - x) <= Rabs (x2 - x1) + Rabs (x1 - x)D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:forall x3 : R, disc x x0 x3 -> V0 x3x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D yH9:0 < x0 - Rabs (x - x1)del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posrealx2:RH10:disc x1 del x2Rabs (x2 - x1) + Rabs (x1 - x) < x0D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y0 < x0 - Rabs (x - x1)D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:forall x3 : R, disc x x0 x3 -> V0 x3x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D yH9:0 < x0 - Rabs (x - x1)del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posrealx2:RH10:disc x1 del x2Rabs (x2 - x1) + Rabs (x1 - x) < x0D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y0 < x0 - Rabs (x - x1)D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:forall x3 : R, disc x x0 x3 -> V0 x3x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D yH9:0 < x0 - Rabs (x - x1)del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posrealx2:RH10:disc x1 del x2Rabs (x2 - x1) + Rabs (x1 - x) < del + Rabs (x1 - x)D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:forall x3 : R, disc x x0 x3 -> V0 x3x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D yH9:0 < x0 - Rabs (x - x1)del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posrealx2:RH10:disc x1 del x2del + Rabs (x1 - x) = x0D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y0 < x0 - Rabs (x - x1)D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:forall x3 : R, disc x x0 x3 -> V0 x3x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D yH9:0 < x0 - Rabs (x - x1)del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posrealx2:RH10:disc x1 del x2del + Rabs (x1 - x) = x0D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y0 < 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).D:R -> Propx:RH:~ (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) -> PropH0:exists n : R -> Prop, ~ P nV0:R -> PropH1:~ (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) V0H4:~ (exists y : R, intersection_domain V0 D y)x0:posrealH5:included (disc x x0) V0x1:RH6:disc x x0 x1H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D yH8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y0 < x0 - Rabs (x - x1)forall D : R -> Prop, open_set D <-> D =_D interior Dforall D : R -> Prop, open_set D <-> D =_D interior DD:R -> Propopen_set D -> D =_D interior DD:R -> PropD =_D interior D -> open_set DD:R -> PropH:open_set Dincluded D (interior D)D:R -> PropH:open_set Dincluded (interior D) DD:R -> PropD =_D interior D -> open_set DD:R -> PropH:open_set Dincluded (interior D) DD:R -> PropD =_D interior D -> open_set Dintro; 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.D:R -> PropD =_D interior D -> open_set Dforall D : R -> Prop, closed_set D <-> D =_D adherence Dforall D : R -> Prop, closed_set D <-> D =_D adherence DD:R -> Propclosed_set D -> D =_D adherence DD:R -> PropD =_D adherence D -> closed_set DD:R -> PropH:closed_set Dincluded D (adherence D)D:R -> PropH:closed_set Dincluded (adherence D) DD:R -> PropD =_D adherence D -> closed_set DD:R -> PropH:closed_set Dincluded (adherence D) DD:R -> PropD =_D adherence D -> closed_set DD:R -> PropD =_D adherence D -> closed_set DD:R -> PropH:(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)) x0x:RH1:complementary D xcomplementary (adherence D) xD:R -> PropH:(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)) x0x:RH1:complementary D xH2:complementary (adherence D) xneighbourhood (complementary D) xassert (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.D:R -> PropH:(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)) x0x:RH1:complementary D xH2:complementary (adherence D) xneighbourhood (complementary D) xforall (D1 D2 : R -> Prop) (x : R), included D1 D2 -> neighbourhood D1 x -> neighbourhood D2 xunfold 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) (x : R), included D1 D2 -> neighbourhood D1 x -> neighbourhood D2 xforall 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 -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D1 xneighbourhood (union_domain D1 D2) xD1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D2 xneighbourhood (union_domain D1 D2) xD1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D1 xincluded D1 (union_domain D1 D2)D1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D1 xneighbourhood D1 xD1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D2 xneighbourhood (union_domain D1 D2) xD1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D1 xneighbourhood D1 xD1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D2 xneighbourhood (union_domain D1 D2) xD1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D2 xneighbourhood (union_domain D1 D2) xD1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D2 xincluded D2 (union_domain D1 D2)D1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D2 xneighbourhood D2 xapply H0; assumption. Qed.D1, D2:R -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x \/ D2 xH2:D2 xneighbourhood D2 xforall 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 -> PropH:forall x0 : R, D1 x0 -> neighbourhood D1 x0H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0x:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xneighbourhood (intersection_domain D1 D2) xD1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2 -> neighbourhood (fun c : R => D1 c /\ D2 c) xD1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D2H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealneighbourhood (fun c : R => D1 c /\ D2 c) xD1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < delD1 x0 /\ D2 x0D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < delD1 x0D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < delD2 x0D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < delRabs (x0 - x) < delD1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < deldel <= del1D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < delD2 x0D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < deldel <= del1D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < delD2 x0D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < delD2 x0D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < delRabs (x0 - x) < delD1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < deldel <= del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1del2:posrealH0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1H6:0 < Rmin del1 del2del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posrealx0:RH7:Rabs (x0 - x) < deldel <= del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D20 < Rmin del1 del2D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D2r:del1 <= del20 < del1D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D2n:~ del1 <= del20 < del2apply (cond_pos del2). Qed.D1, D2:R -> Propx:RH1:D1 x /\ D2 xH2:D1 xH3:D2 xH4:exists delta : posreal, included (disc x delta) D1H5:exists delta : posreal, included (disc x delta) D2del1:posrealH:included (disc x del1) D1del2:posrealH0:included (disc x del2) D2n:~ del1 <= del20 < del2open_set (fun _ : R => False)unfold open_set; intros; elim H. Qed.open_set (fun _ : R => False)open_set (fun _ : R => True)open_set (fun _ : R => True)exists (mkposreal 1 Rlt_0_1); unfold included; intros; trivial. Qed.x:RH:Trueexists delta : posreal, included (disc x delta) (fun _ : R => True)forall (x : R) (del : posreal), open_set (disc x del)forall (x : R) (del : posreal), open_set (disc x del)x:Rdel:posrealH:open_set (disc x del) <-> disc x del =_D interior (disc x del)open_set (disc x del)x:Rdel:posrealH: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:Rdel:posrealH: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:Rdel:posrealH: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:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0) -> neighbourhood (fun y : R => Rabs (y - x) < del) x0x:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0)x:Rdel:posrealH: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:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealneighbourhood (fun y : R => Rabs (y - x) < del) x0x:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0)x:Rdel:posrealH: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:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealx1:RH4:disc x0 del2 x1Rabs (x1 - x) < delx:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0)x:Rdel:posrealH: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:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealx1:RH4:disc x0 del2 x1Rabs (x1 - x) <= Rabs (x1 - x0) + Rabs (x0 - x)x:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealx1:RH4:disc x0 del2 x1Rabs (x1 - x0) + Rabs (x0 - x) < delx:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0)x:Rdel:posrealH: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:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealx1:RH4:disc x0 del2 x1Rabs (x1 - x0) + Rabs (x0 - x) < delx:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0)x:Rdel:posrealH: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:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealx1:RH4:disc x0 del2 x1Rabs (x1 - x0) + Rabs (x0 - x) < del2 + Rabs (x0 - x)x:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealx1:RH4:disc x0 del2 x1del2 + Rabs (x0 - x) = delx:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0)x:Rdel:posrealH: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:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealx1:RH4:disc x0 del2 x1Rabs (x1 - x0) < del2x:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealx1:RH4:disc x0 del2 x1del2 + Rabs (x0 - x) = delx:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0)x:Rdel:posrealH: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:Rdel:posrealH: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:RH2:Rabs (x0 - x) < delH3:0 < del - Rabs (x - x0)del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posrealx1:RH4:disc x0 del2 x1del2 + Rabs (x0 - x) = delx:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0)x:Rdel:posrealH: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:Rdel:posrealH: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:RH2:Rabs (x0 - x) < del0 < del - Rabs (x - x0)x:Rdel:posrealH: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.x:Rdel:posrealH: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)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 -> Rx:Rcontinuity_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 -> Rx: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 xf:R -> Rx:RH:continuity_pt f xW:R -> PropH0:exists delta : posreal, included (disc (f x) delta) Wexists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))f:R -> Rx: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 xf:R -> Rx:RH:continuity_pt f xW:R -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:included (disc (f x) del1) Wexists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:included (disc (f x) del1) Wexists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:included (disc (f x) del1) WH2: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 -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:included (disc (f x) del1) WH2: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:RH3: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 -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:included (disc (f x) del1) WH2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:included (disc (f x) del1) WH2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1neighbourhood (disc x {| pos := del2; cond_pos := H4 |}) x /\ (forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y))f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1neighbourhood (disc x {| pos := del2; cond_pos := H4 |}) xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y)f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1exists delta : posreal, included (fun y : R => Rabs (y - x) < delta) (fun y : R => Rabs (y - x) < {| pos := del2; cond_pos := H4 |})f:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y)f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1included (fun y : R => Rabs (y - x) < {| pos := del2; cond_pos := H4 |}) (fun y : R => Rabs (y - x) < {| pos := del2; cond_pos := H4 |})f:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y)f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y)f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y = xRabs (f y - f x) < del1f:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xRabs (f y - f x) < del1f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xRabs (f y - f x) < del1f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xD_x no_cond x yf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xRabs (y - x) < del2f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xTruef:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xx <> yf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xRabs (y - x) < del2f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xx <> yf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xRabs (y - x) < del2f:R -> Rx: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 xf:R -> Rx:RH: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 -> PropH0:exists delta : posreal, included (disc (f x) delta) Wdel1:posrealH1:forall x0 : R, disc (f x) del1 x0 -> W x0H2: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:RH3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)H4:del2 > 0H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1y:RH6:disc x {| pos := del2; cond_pos := H4 |} yH7:y <> xRabs (y - x) < del2f:R -> Rx: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 xf:R -> Rx: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 xf:R -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0exists 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 -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> PropH4: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) DH6:forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y)del1:posrealH7:included (disc x del1) Dexists 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 -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> PropH4: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) DH6:forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y)del1:posrealH7:included (disc x del1) Ddel1 > 0f:R -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> PropH4: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) DH6:forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y)del1:posrealH7:included (disc x del1) Dforall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < del1 -> dist R_met (f x0) (f x) < epsf:R -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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 -> PropH4: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) DH6:forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y)del1:posrealH7:included (disc x del1) Dforall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < del1 -> dist R_met (f x0) (f x) < epsf:R -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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). (**********)f:R -> Rx:RH:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))eps:RH0:eps > 0H1: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)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) (D : R -> Prop), continuity f -> 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))forall f : R -> R, continuity f <-> (forall D : R -> Prop, open_set D -> open_set (image_rec f D))f:R -> Rcontinuity 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 ff:R -> R(forall D : R -> Prop, open_set D -> open_set (image_rec f D)) -> continuity ff:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_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 -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_set (disc (f x) {| pos := eps; cond_pos := H0 |})f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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 -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_set (disc (f x) {| pos := eps; cond_pos := H0 |})f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0disc (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 -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0disc (f x) {| pos := eps; cond_pos := H0 |} (f x)f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_set (disc (f x) {| pos := eps; cond_pos := H0 |})f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0H3: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)) xexists 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 -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0disc (f x) {| pos := eps; cond_pos := H0 |} (f x)f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_set (disc (f x) {| pos := eps; cond_pos := H0 |})f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0H3: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:posrealH5: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 -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0disc (f x) {| pos := eps; cond_pos := H0 |} (f x)f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_set (disc (f x) {| pos := eps; cond_pos := H0 |})f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0H3: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:posrealH5:included (disc x del) (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0))del > 0f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0H3: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:posrealH5: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) < epsf:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0disc (f x) {| pos := eps; cond_pos := H0 |} (f x)f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_set (disc (f x) {| pos := eps; cond_pos := H0 |})f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0H3: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:posrealH5: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) < epsf:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0disc (f x) {| pos := eps; cond_pos := H0 |} (f x)f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_set (disc (f x) {| pos := eps; cond_pos := H0 |})f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0H1: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)) x0disc (f x) {| pos := eps; cond_pos := H0 |} (f x)f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_set (disc (f x) {| pos := eps; cond_pos := H0 |})apply disc_P1. Qed. (**********)f:R -> RH:forall D : R -> Prop, open_set D -> open_set (image_rec f D)x, eps:RH0:eps > 0open_set (disc (f x) {| pos := eps; cond_pos := H0 |})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:RHsep:x <> yD:=Rabs (x - y):Rexists V W : R -> Prop, neighbourhood V x /\ neighbourhood W y /\ ~ (exists y0 : R, intersection_domain V W y0)x, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2 -> exists V W : R -> Prop, neighbourhood V x /\ neighbourhood W y /\ ~ (exists y0 : R, intersection_domain V W y0)x, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2exists 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:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2neighbourhood (disc x {| pos := D / 2; cond_pos := H |}) xx, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2neighbourhood (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:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2neighbourhood (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:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2neighbourhood (disc y {| pos := D / 2; cond_pos := H |}) yx, y:RHsep:x <> yD:=Rabs (x - y):RH: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:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH: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:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0Falsex, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0D < D -> Falsex, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0D < Dx, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0D < Dx, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0Rabs (x - y) <= Rabs (x - x0) + Rabs (x0 - y)x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0Rabs (x - x0) + Rabs (x0 - y) < Dx, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0Rabs (x - x0) + Rabs (x0 - y) < Dx, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0Rabs (x - x0) < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0Rabs (x0 - y) < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):RH:0 < D / 2H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0x0:RH1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0H2:disc x {| pos := D / 2; cond_pos := H |} x0H3:disc y {| pos := D / 2; cond_pos := H |} x0Rabs (x0 - y) < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):R0 < D / 2x, y:RHsep:x <> yD:=Rabs (x - y):R0 < Dx, y:RHsep:x <> yD:=Rabs (x - y):R0 < / 2apply 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.x, y:RHsep:x <> yD:=Rabs (x - y):R0 < / 2forall (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 xforall (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 xf0:familyD:R -> Propx:RH:exists y : R, (fun z1 z2 : R => f0 z1 z2 /\ D z1) x yx0:RH0:f0 x x0 /\ D xH1:f0 x x0H2:D xind f0 xf0:familyD:R -> Propx:RH:exists y : R, (fun z1 z2 : R => f0 z1 z2 /\ D z1) x yx0:RH0:f0 x x0 /\ D xH1:f0 x x0H2:D xD xassumption. 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). (**********)f0:familyD:R -> Propx:RH:exists y : R, (fun z1 z2 : R => f0 z1 z2 /\ D z1) x yx0:RH0:f0 x x0 /\ D xH1:f0 x x0H2:D xD xforall (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:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xopen_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:D xopen_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:~ D xopen_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1: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:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:D xopen_set (f0 x) -> open_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:~ D xopen_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:D xopen_set (f0 x) -> open_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:~ D xopen_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x3 : R, open_set (f0 x3)x:RH0:D x \/ ~ D xH1:D xH2:forall x3 : R, f0 x x3 -> exists delta : posreal, included (disc x3 delta) (f0 x)x0:RH3:f0 x x0 /\ D xH4:f0 x x0H5:D xH6:exists delta : posreal, included (disc x0 delta) (f0 x)x1:posrealH7:included (disc x0 x1) (f0 x)x2:RH8:disc x0 x1 x2f0 x x2f0:familyD:R -> PropH:forall x3 : R, open_set (f0 x3)x:RH0:D x \/ ~ D xH1:D xH2:forall x3 : R, f0 x x3 -> exists delta : posreal, included (disc x3 delta) (f0 x)x0:RH3:f0 x x0 /\ D xH4:f0 x x0H5:D xH6:exists delta : posreal, included (disc x0 delta) (f0 x)x1:posrealH7:included (disc x0 x1) (f0 x)x2:RH8:disc x0 x1 x2D xf0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:~ D xopen_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x3 : R, open_set (f0 x3)x:RH0:D x \/ ~ D xH1:D xH2:forall x3 : R, f0 x x3 -> exists delta : posreal, included (disc x3 delta) (f0 x)x0:RH3:f0 x x0 /\ D xH4:f0 x x0H5:D xH6:exists delta : posreal, included (disc x0 delta) (f0 x)x1:posrealH7:included (disc x0 x1) (f0 x)x2:RH8:disc x0 x1 x2D xf0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:~ D xopen_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:~ D xopen_set (fun y : R => f0 x y /\ D x)f0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:~ 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:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:~ D xopen_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)).f0:familyD:R -> PropH:forall x0 : R, open_set (f0 x0)x:RH0:D x \/ ~ D xH1:~ D xopen_set (fun _ : R => False) -> open_set (fun y : R => f0 x y /\ D x)forall D1 D2 : R -> Prop, open_set D1 -> D1 =_D D2 -> open_set D2forall D1 D2 : R -> Prop, open_set D1 -> D1 =_D D2 -> open_set D2D1, D2:R -> PropH:forall x0 : R, D1 x0 -> exists delta : posreal, included (disc x0 delta) D1H0:D1 =_D D2x:RH1:D2 xexists delta : posreal, included (disc x delta) D2D1, D2:R -> PropH:forall x0 : R, D1 x0 -> exists delta : posreal, included (disc x0 delta) D1H0:included D1 D2 /\ included D2 D1x:RH1:D2 xH2:included D1 D2H3:included D2 D1exists delta : posreal, included (disc x delta) D2D1, D2:R -> PropH:forall x0 : R, D1 x0 -> exists delta : posreal, included (disc x0 delta) D1H0:included D1 D2 /\ included D2 D1x:RH1:D2 xH2:included D1 D2H3:included D2 D1H4:exists delta : posreal, included (disc x delta) D1exists delta : posreal, included (disc x delta) D2exists x0; apply included_trans with D1; assumption. Qed. (**********)D1, D2:R -> PropH:forall x1 : R, D1 x1 -> exists delta : posreal, included (disc x1 delta) D1H0:included D1 D2 /\ included D2 D1x:RH1:D2 xH2:included D1 D2H3:included D2 D1H4:exists delta : posreal, included (disc x delta) D1x0:posrealH5:included (disc x x0) D1exists delta : posreal, included (disc x delta) D2forall X : R -> Prop, compact X -> bounded Xforall X : R -> Prop, compact X -> bounded XX:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truebounded XX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0 -> bounded XX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4: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 l0l:RlistH7:forall x : R, ind (subfamily f0 D') x <-> In x lr:=MaxRlist l:Rexists m M : R, forall x : R, X x -> m <= x <= MX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4: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 l0l:RlistH7:forall x0 : R, ind (subfamily f0 D') x0 <-> In x0 lr:=MaxRlist l:Rx:RH8:X x- r <= x <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0 -> - r <= x <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0- r <= x <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 l- r <= xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= r -> - r <= xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lH17:x0 <= rRabs x < r -> - r <= xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lH17:x0 <= rRabs x < rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lH17:x0 <= rRabs x < rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= r -> x <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lH17:x0 <= rx <= Rabs xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lH17:x0 <= rRabs x <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lH17:x0 <= rRabs x <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lH17:x0 <= rRabs x <= x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lH17:x0 <= rx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lH17:x0 <= rx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:Rabs x < x0H12:D' x0H13:intersection_domain D D' x0 -> In x0 lH14:In x0 l -> intersection_domain D D' x0H15:intersection_domain D D' x0H16:In x0 lx0 <= rX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x1 y : R => Rabs y < x1:R -> R -> PropH0:forall x1 : R, (exists y : R, g x1 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H2:covering_open_set X f0H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)D':R -> PropH4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0l:RlistH7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 lr:=MaxRlist l:Rx:RH8:X xH9:exists y : R, subfamily f0 D' y xx0:RH10:g x0 x /\ D' x0H11:g x0 xH12:D' x0H13:intersection_domain D D' x0 <-> In x0 lintersection_domain D D' x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)family_open_set f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x y : R => Rabs y < x:R -> R -> PropH0:forall x : R, (exists y : R, g x y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)family_open_set f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 < xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 < xopen_set (disc 0 {| pos := x; cond_pos := H2 |})X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 < xdisc 0 {| pos := x; cond_pos := H2 |} =_D f0 xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 < xdisc 0 {| pos := x; cond_pos := H2 |} =_D f0 xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 < xincluded (fun y : R => Rabs (y - 0) < {| pos := x; cond_pos := H2 |}) (fun y : R => Rabs y < x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 < xincluded (fun y : R => Rabs y < x) (fun y : R => Rabs (y - 0) < {| pos := x; cond_pos := H2 |})X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 < xincluded (fun y : R => Rabs y < x) (fun y : R => Rabs (y - 0) < {| pos := x; cond_pos := H2 |})X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xopen_set (fun _ : R => False)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > x(fun _ : R => False) =_D f0 xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > x(fun _ : R => False) =_D f0 xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xincluded (fun _ : R => False) (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xincluded (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. (**********)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)D:=fun _ : R => True:R -> Propg:=fun x0 y : R => Rabs y < x0:R -> R -> PropH0:forall x0 : R, (exists y : R, g x0 y) -> Truef0:={| ind := D; f := g; cond_fam := H0 |}:familyH1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH2:0 = x \/ 0 > xincluded (f0 x) (fun _ : R => False)forall X : R -> Prop, compact X -> closed_set Xforall X : R -> Prop, compact X -> closed_set XX:R -> PropH:compact XX =_D adherence XX:R -> PropH:compact Xincluded X (adherence X)X:R -> PropH:compact Xincluded (adherence X) XX:R -> PropH:compact Xincluded (adherence X) XX:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:X xX xX:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xX xX:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xX xX:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X x(forall y : R, X y -> 0 < Rabs (y - x) / 2) -> X xX:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=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 xX:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0 -> D xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering_finite X (subfamily f0 D')D xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alp -> D xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yneighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> D xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yneighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0alp <= Rabs (y0 - x) / 2 -> D xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0alp <= Rabs (y0 - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yneighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0H17:alp <= Rabs (y0 - x) / 2H18:Rabs (y - x) < Rabs (y0 - x) / 2Rabs (y0 - x) < Rabs (y0 - x) -> D xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0H17:alp <= Rabs (y0 - x) / 2H18:Rabs (y - x) < Rabs (y0 - x) / 2Rabs (y0 - x) < Rabs (y0 - x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0alp <= Rabs (y0 - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yneighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0H17:alp <= Rabs (y0 - x) / 2H18:Rabs (y - x) < Rabs (y0 - x) / 2Rabs (y0 - x) < Rabs (y0 - x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0alp <= Rabs (y0 - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yneighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0H17:alp <= Rabs (y0 - x) / 2H18:Rabs (y - x) < Rabs (y0 - x) / 2Rabs (y0 - x) <= Rabs (y0 - y) + Rabs (y - x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0H17:alp <= Rabs (y0 - x) / 2H18:Rabs (y - x) < Rabs (y0 - x) / 2Rabs (y0 - y) + Rabs (y - x) < Rabs (y0 - x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0alp <= Rabs (y0 - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yneighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0H17:alp <= Rabs (y0 - x) / 2H18:Rabs (y - x) < Rabs (y0 - x) / 2Rabs (y0 - y) + Rabs (y - x) < Rabs (y0 - x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0alp <= Rabs (y0 - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yneighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1H1:~ X xH2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2D:=X:R -> Propg:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> PropH3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1l:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xy:RH12:Rabs (y - x) < alpH13:X yy0:RH15:D' y0H14:Rabs (y0 - y) < Rabs (y0 - x) / 2H16:D y0alp <= Rabs (y0 - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yneighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yneighbourhood (disc x {| pos := alp; cond_pos := H9 |}) xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):RH9:0 < alpH10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X yH11:forall x0 : R, disc x {| pos := alp; cond_pos := H9 |} x0 -> neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x0disc x {| pos := alp; cond_pos := H9 |} xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)H5:covering_open_set X f0D':R -> PropH6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' yl:RlistH8:forall x0 : R, intersection_domain D D' x0 <-> In x0 lalp:=MinRlist (AbsList l x):R0 < alpX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering_open_set X f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)covering X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)family_open_set f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:X x0Rabs (x0 - x0) < Rabs (x0 - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:X x0D x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)family_open_set f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:X x0D x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)family_open_set f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)family_open_set f0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0open_set (disc x0 {| pos := Rabs (x0 - x) / 2; cond_pos := H2 x0 H5 |})X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0disc 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 -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0disc 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 -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0included (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 -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0included (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 -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x2 : R, (exists y : R, g x2 y) -> D x2f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0x1:RH6:Rabs (x1 - x0) < Rabs (x0 - x) / 2Rabs (x0 - x1) < Rabs (x0 - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x2 : R, (exists y : R, g x2 y) -> D x2f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0x1:RH6:Rabs (x1 - x0) < Rabs (x0 - x) / 2D x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0included (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 -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x2 : R, (exists y : R, g x2 y) -> D x2f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0x1:RH6:Rabs (x1 - x0) < Rabs (x0 - x) / 2D x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0included (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 -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:D x0included (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 -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0open_set (fun _ : R => False)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0(fun _ : R => False) =_D (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0(fun _ : R => False) =_D (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0included (fun _ : R => False) (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0included (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0) (fun _ : R => False)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> PropH3:forall x1 : R, (exists y : R, g x1 y) -> D x1f0:={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x0:RH5:~ D x0included (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0) (fun _ : R => False)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xH2:forall y : R, X y -> 0 < Rabs (y - x) / 2D:=X:R -> Propg:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Propforall x0 : R, (exists y : R, g x0 y) -> D x0X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X yH1:~ X xforall y : R, X y -> 0 < Rabs (y - x) / 2X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y0 : R, intersection_domain V X y0H1:~ X xy:RH2:X y0 < Rabs (y - x)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y0 : R, intersection_domain V X y0H1:~ X xy:RH2:X y0 < / 2apply Rinv_0_lt_compat; prove_sup0. Qed. (**********)X:R -> PropH:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)x:RH0:forall V : R -> Prop, neighbourhood V x -> exists y0 : R, intersection_domain V X y0H1:~ X xy:RH2:X y0 < / 2compact (fun _ : R => False)compact (fun _ : R => False)f0:familyH:covering_open_set (fun _ : R => False) f0covering (fun _ : R => False) (subfamily f0 (fun _ : R => False))f0:familyH:covering_open_set (fun _ : R => False) f0family_finite (subfamily f0 (fun _ : R => False))f0:familyH:covering_open_set (fun _ : R => False) f0family_finite (subfamily f0 (fun _ : R => False))f0:familyH:covering_open_set (fun _ : R => False) f0x:Rind (subfamily f0 (fun _ : R => False)) x <-> In x nilf0:familyH:covering_open_set (fun _ : R => False) f0x:Rind (subfamily f0 (fun _ : R => False)) x -> In x nilf0:familyH:covering_open_set (fun _ : R => False) f0x:RIn x nil -> ind (subfamily f0 (fun _ : R => False)) xf0:familyH:covering_open_set (fun _ : R => False) f0x:RH0:ind f0 x /\ Falseind f0 x -> False -> Falsef0:familyH:covering_open_set (fun _ : R => False) f0x:RIn x nil -> ind (subfamily f0 (fun _ : R => False)) xsimpl; intro; elim H0. Qed.f0:familyH:covering_open_set (fun _ : R => False) f0x:RIn x nil -> ind (subfamily f0 (fun _ : R => False)) xforall X1 X2 : R -> Prop, compact X1 -> X1 =_D X2 -> compact X2forall X1 X2 : R -> Prop, compact X1 -> X1 =_D X2 -> compact X2X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)f0:familyH1:covering_open_set X2 f0H0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xcovering_open_set X1 f0X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)f0:familyH1:covering_open_set X2 f0H0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH3:covering_open_set X1 f0exists D : R -> Prop, covering_finite X2 (subfamily f0 D)X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)f0:familyH0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH1:covering X2 f0H3:family_open_set f0covering X1 f0X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)f0:familyH0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH1:covering X2 f0H3:family_open_set f0family_open_set f0X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)f0:familyH1:covering_open_set X2 f0H0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH3:covering_open_set X1 f0exists D : R -> Prop, covering_finite X2 (subfamily f0 D)X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)f0:familyH0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH1:covering X2 f0H3:family_open_set f0family_open_set f0X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)f0:familyH1:covering_open_set X2 f0H0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH3:covering_open_set X1 f0exists D : R -> Prop, covering_finite X2 (subfamily f0 D)X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)f0:familyH1:covering_open_set X2 f0H0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH3:covering_open_set X1 f0exists D : R -> Prop, covering_finite X2 (subfamily f0 D)X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D0 : R -> Prop, covering_finite X1 (subfamily f1 D0)f0:familyH1:covering_open_set X2 f0H0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH3:covering_open_set X1 f0D:R -> PropH4: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 -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D0 : R -> Prop, covering_finite X1 (subfamily f1 D0)f0:familyH1:covering_open_set X2 f0H0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH3:covering_open_set X1 f0D:R -> PropH4: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.X1, X2:R -> PropH:forall f1 : family, covering_open_set X1 f1 -> exists D0 : R -> Prop, covering_finite X1 (subfamily f1 D0)f0:familyH1:covering_open_set X2 f0H0:forall x : R, X1 x -> X2 xH2:forall x : R, X2 x -> X1 xH3:covering_open_set X1 f0D:R -> PropH4: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)
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:RHle:a <= bcompact (fun c : R => a <= c <= b)a, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3: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:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 mexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propcovering (fun c : R => a <= c <= b) (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHle':x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHle':x0 <= xH15:a <= x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHle':x0 <= xa <= x0 <= xa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHle':x0 <= xa <= x0 <= xa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xf0 y0 x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xb - x0 < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xb - x0 >= 0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xb - x0 < b - xa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xb - x < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xb - x0 >= 0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xb - x < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xb - x0 >= 0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xb - x0 >= 0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= bHnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:b - eps < xDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propexists l : Rlist, forall x0 : R, ind (subfamily f0 Db) x0 <-> In x0 la, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:Rind (subfamily f0 Db) x0 -> In x0 (cons y0 l)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ Db x0H16:x0 = y0In x0 (cons y0 l)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ Db x0H16:x0 <> y0In x0 (cons y0 l)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ Db x0H16:x0 <> y0In x0 (cons y0 l)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ Db x0H16:x0 <> y0ind (subfamily f0 Dx) x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)H16:x0 <> y0H7:ind f0 x0H11:Dx x0ind f0 x0 /\ Dx x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)H16:x0 <> y0H7:ind f0 x0H11:x0 = y0ind f0 x0 /\ Dx x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)H16:x0 <> y0H7:ind f0 x0H11:x0 = y0ind f0 x0 /\ Dx x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:x0 = y0ind f0 x0 /\ Db x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lind f0 x0 /\ Db x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:x0 = y0ind f0 x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:x0 = y0Db x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lind f0 x0 /\ Db x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:x0 = y0Db x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lind f0 x0 /\ Db x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lind f0 x0 /\ Db x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 l(ind (subfamily f0 Dx) x0 -> In x0 l) -> (In x0 l -> ind (subfamily f0 Dx) x0) -> ind f0 x0 /\ Db x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0H17:ind f0 x0 /\ Dx x0ind f0 x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0H17:ind f0 x0 /\ Dx x0Db x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0H4:a <= b <= bH3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)y0:RH6:f0 y0 beps:posrealH8:included (disc b eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:b - eps < xDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0H17:ind f0 x0 /\ Dx x0Db x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mexists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mH17:m < m'exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mm < m'a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mm < m'a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mHle':m + eps / 2 <= bm < m + eps / 2a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mHnle':~ m + eps / 2 <= bm < ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mHnle':~ m + eps / 2 <= bm < ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H3:m < by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mHnle':~ m + eps / 2 <= bm < ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H3:m = by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mHnle':~ m + eps / 2 <= bm < ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH14:forall x0 : R, A x0 -> x0 <= mH15:forall b0 : R, is_upper_bound A b0 -> m <= b0H3:m = by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RH7:A m'H16:m' <= mHnle':~ m + eps / 2 <= bm < ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RA m'a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Ra <= m' <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Ra <= m'a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rm' <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Ra <= ma, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rm <= m'a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rm' <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rm <= m'a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rm' <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rr:m + eps / 2 <= bm <= m + eps / 2a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rn:~ m + eps / 2 <= bm <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rm' <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rn:~ m + eps / 2 <= bm <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rm' <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= mH7:m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rn:~ m + eps / 2 <= bm <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rm' <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rm' <= ba, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:Rexists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propcovering (fun c : R => a <= c <= m') (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hle':x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hle':x0 <= xH15:a <= x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hle':x0 <= xa <= x0 <= xa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hle':x0 <= xa <= x0 <= xa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xexists y : R, subfamily f0 Db y x0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xf0 y0 x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0- (x0 - m) < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0x0 - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0m - x0 < m - xa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0m - x < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0x0 - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0m - x < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0x0 - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0m - eps < x - eps + epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0m - eps = x - eps + (m - x)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0x0 - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0m - eps < xa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0x = x - eps + epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0m - eps = x - eps + (m - x)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0x0 - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0x = x - eps + epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0m - eps = x - eps + (m - x)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0x0 - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHlt:x0 - m < 0m - eps = x - eps + (m - x)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0x0 - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0x0 - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0x0 - m <= m' - ma, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' - m < epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' < m + epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' = m + (m' - m)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' <= m + eps / 2a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m + eps / 2 < m + epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' = m + (m' - m)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m + eps / 2 < m + epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' = m + (m' - m)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 00 < 2a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 02 * (eps / 2) < 2 * epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' = m + (m' - m)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 02 * (eps / 2) < 2 * epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' = m + (m' - m)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 01 * eps < 2 * epsa, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 02 <> 0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' = m + (m' - m)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 02 <> 0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' = m + (m' - m)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xHge:x0 - m >= 0m' = m + (m' - m)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH14:a <= x0H18:x0 <= m'Hnle':~ x0 <= xDb y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)H13:family_finite (subfamily f0 Dx)Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Propfamily_finite (subfamily f0 Db)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:Rind (subfamily f0 Db) x0 -> In x0 (cons y0 l)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ Db x0Heq:x0 = y0In x0 (cons y0 l)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ Db x0Hneq:x0 <> y0In x0 (cons y0 l)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ Db x0Hneq:x0 <> y0In x0 (cons y0 l)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)Hneq:x0 <> y0H7:ind f0 x0H16:Dx x0ind f0 x0 /\ Dx x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)Hneq:x0 <> y0H7:ind f0 x0H16:x0 = y0ind f0 x0 /\ Dx x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:Rlistx0:RH13:ind (subfamily f0 Dx) x0 -> In x0 lH15:In x0 l -> ind (subfamily f0 Dx) x0Hltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> PropH14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)Hneq:x0 <> y0H7:ind f0 x0H16:x0 = y0ind f0 x0 /\ Dx x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RIn x0 (cons y0 l) -> ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:x0 = y0ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:x0 = y0ind f0 x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:x0 = y0Db x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:x0 = y0Db x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0H17:ind (subfamily f0 Dx) x0ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0H17:intersection_domain (ind f0) Dx x0ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0H17:ind f0 x0 /\ Dx x0ind (subfamily f0 Db) x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0H17:ind f0 x0 /\ Dx x0ind f0 x0a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0H17:ind f0 x0 /\ Dx x0Db x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)x:RH9:a <= x <= bDx:R -> PropH12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)l:RlistH13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 lHltx:m - eps < xH11:m <> bm':=Rmin (m + eps / 2) b:RDb:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Propx0:RH15:In x0 lH16:In x0 l -> ind (subfamily f0 Dx) x0H17:ind f0 x0 /\ Dx x0Db x0a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:exists x : R, A x /\ m - eps < x <= mexists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0is_upper_bound A (m - eps) -> exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0is_upper_bound A (m - eps)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0H12:is_upper_bound A (m - eps)H13:m <= m - epsm - eps < m -> exists x : R, A x /\ m - eps < x <= ma, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0H12:is_upper_bound A (m - eps)H13:m <= m - epsm - eps < ma, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0is_upper_bound A (m - eps)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0H12:is_upper_bound A (m - eps)H13:m <= m - epsm - eps < ma, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0is_upper_bound A (m - eps)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x : R, A x /\ m - eps < x <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0is_upper_bound A (m - eps)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ A xx <= m - epsa, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mx <= m - epsa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mx <= m - epsa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mH16:~ m - eps < xx <= m - epsa, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mH16:~ x <= mx <= m - epsa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mH16:~ m - eps < xH17:x <= m - epsx <= m - epsa, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mH16:~ m - eps < xH17:~ x <= m - epsx <= m - epsa, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mH16:~ x <= mx <= m - epsa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mH16:~ m - eps < xH17:~ x <= m - epsx <= m - epsa, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mH16:~ x <= mx <= m - epsa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5: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 -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)H4:a <= m <= by0:RH6:f0 y0 meps:posrealH8:included (disc m eps) (f0 y0)H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)H10:is_upper_bound A mH11:forall b0 : R, is_upper_bound A b0 -> m <= b0P:=fun n : R => A n /\ m - eps < n <= m:R -> PropH12:forall n : R, ~ (A n /\ m - eps < n <= m)x:RH13:A xH14:~ A x \/ ~ m - eps < x <= mH15:~ m - eps < x <= mH16:~ x <= mx <= m - epsa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:is_upper_bound A mH4:forall b0 : R, is_upper_bound A b0 -> m <= b0a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:forall x : R, A x -> x <= mH4:forall b0 : R, is_upper_bound A b0 -> m <= b0a <= m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:forall x : R, A x -> x <= mH4:forall b0 : R, is_upper_bound A b0 -> m <= b0a <= ma, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:forall x : R, A x -> x <= mH4:forall b0 : R, is_upper_bound A b0 -> m <= b0m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:forall x : R, A x -> x <= mH4:forall b0 : R, is_upper_bound A b0 -> m <= b0m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound AH2:exists a0 : R, A a0m:RH3:forall x : R, A x -> x <= mH4:forall b0 : R, is_upper_bound A b0 -> m <= b0m <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A aH1:bound Aexists a0 : R, A a0a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH0:A abound Aa, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropA aa, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propexists D : R -> Prop, covering_finite (fun c : R => a <= c <= a) (subfamily f0 D)a, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:covering (fun c : R => a <= c <= b) f0H5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propexists D : R -> Prop, covering_finite (fun c : R => a <= c <= a) (subfamily f0 D)a, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= b -> exists D : R -> Prop, covering_finite (fun c : R => a <= c <= a) (subfamily f0 D)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x : R => x = y0:R -> Propcovering (fun c : R => a <= c <= a) (subfamily f0 D')a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x : R => x = y0:R -> Propfamily_finite (subfamily f0 D')a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH3:a <= x <= ax = a -> exists y : R, f0 y x /\ D' ya, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH3:a <= x <= ax = aa, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x : R => x = y0:R -> Propfamily_finite (subfamily f0 D')a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH3:a <= x <= aH4:x = af0 y0 xa, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH3:a <= x <= aH4:x = aD' y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH3:a <= x <= ax = aa, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x : R => x = y0:R -> Propfamily_finite (subfamily f0 D')a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH3:a <= x <= aH4:x = aD' y0a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH3:a <= x <= ax = aa, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x : R => x = y0:R -> Propfamily_finite (subfamily f0 D')a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH3:a <= x <= ax = aa, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x : R => x = y0:R -> Propfamily_finite (subfamily f0 D')a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x : R => x = y0:R -> Propfamily_finite (subfamily f0 D')a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:Rind (subfamily f0 D') x -> In x (cons y0 nil)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RIn x (cons y0 nil) -> ind (subfamily f0 D') xa, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH3:ind f0 xH4:D' xx = y0 \/ Falsea, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RIn x (cons y0 nil) -> ind (subfamily f0 D') xa, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RIn x (cons y0 nil) -> ind (subfamily f0 D') xa, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0H5:family_open_set f0A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> PropH1:a <= a <= by0:RH2:f0 y0 aD':=fun x0 : R => x0 = y0:R -> Propx:RH4:x = y0ind f0 x /\ D' xa, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHle:a <= bf0:familyH:forall x : R, a <= x <= b -> exists y : R, f0 y xH5:family_open_set f0A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Propa <= a <= ba, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHnle:~ a <= bcompact (fun c : R => a <= c <= b)a, b:RHnle:~ a <= bcompact (fun _ : R => False)a, b:RHnle:~ a <= b(fun _ : R => False) =_D (fun c : R => a <= c <= b)a, b:RHnle:~ a <= b(fun _ : R => False) =_D (fun c : R => a <= c <= b)a, b:RHnle:~ a <= bincluded (fun _ : R => False) (fun c : R => a <= c <= b)a, b:RHnle:~ a <= bincluded (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.a, b:RHnle:~ a <= bincluded (fun c : R => a <= c <= b) (fun _ : R => False)forall X F : R -> Prop, compact X -> closed_set F -> included F X -> compact Fforall X F : R -> Prop, compact X -> closed_set F -> included F X -> compact FX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zexists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propexists D0 : R -> Prop, covering_finite F (subfamily f0 D0)X, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Propexists D0 : R -> Prop, covering_finite F (subfamily f0 D0)X, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propexists D0 : R -> Prop, covering_finite F (subfamily f0 D0)X, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=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 -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f' -> exists D0 : R -> Prop, covering_finite F (subfamily f0 D0)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering_finite X (subfamily f' DX)covering_finite F (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)covering F (subfamily f0 DX) /\ family_finite (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)covering F (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)family_finite (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xexists y : R, subfamily f0 DX y xX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)family_finite (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xy0:RH8:g' y0 xH9:DX y0f0 y0 x /\ DX y0X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)family_finite (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xy0:RH8:g' y0 xH9:DX y0f0 y0 xX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xy0:RH8:g' y0 xH9:DX y0DX y0X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)family_finite (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xy0:RH8:f0 y0 x \/ complementary F x /\ D y0H9:DX y0H10:f0 y0 xf0 y0 xX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xy0:RH8:f0 y0 x \/ complementary F x /\ D y0H9:DX y0H10:complementary F x /\ D y0f0 y0 xX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xy0:RH8:g' y0 xH9:DX y0DX y0X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)family_finite (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xy0:RH8:f0 y0 x \/ complementary F x /\ D y0H9:DX y0H10:complementary F x /\ D y0f0 y0 xX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xy0:RH8:g' y0 xH9:DX y0DX y0X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)family_finite (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0H6:family_finite (subfamily f' DX)x:RH7:F xy0:RH8:g' y0 xH9:DX y0DX y0X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)family_finite (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)H6:family_finite (subfamily f' DX)family_finite (subfamily f0 DX)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)l:RlistH6:forall x0 : R, ind (subfamily f' DX) x0 <-> In x0 lx:RH7:ind (subfamily f' DX) x -> In x lH8:In x l -> ind (subfamily f' DX) xind (subfamily f0 DX) x <-> In x lX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)l:RlistH6:forall x0 : R, ind (subfamily f' DX) x0 <-> In x0 lx:RH7:ind (subfamily f' DX) x -> In x lH8:In x l -> ind (subfamily f' DX) xind (subfamily f0 DX) x -> In x lX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)l:RlistH6:forall x0 : R, ind (subfamily f' DX) x0 <-> In x0 lx:RH7:ind (subfamily f' DX) x -> In x lH8:In x l -> ind (subfamily f' DX) xIn x l -> ind (subfamily f0 DX) xX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH4:covering_open_set X f'DX:R -> PropH5:covering X (subfamily f' DX)l:RlistH6:forall x0 : R, ind (subfamily f' DX) x0 <-> In x0 lx:RH7:ind (subfamily f' DX) x -> In x lH8:In x l -> ind (subfamily f' DX) xIn x l -> ind (subfamily f0 DX) xX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familycovering_open_set X f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0covering X f' /\ family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0covering X f'X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xexists y : R, f' y xX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:F xexists y : R, f' y xX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F xexists y : R, f' y xX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F xexists y : R, f' y xX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F x(exists z : R, D z) -> exists y : R, f' y xX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F xexists z : R, D zX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x1 : R, F x1 -> exists y : R, f0 y x1H4:family_open_set f0x:RH5:X xH6:~ F xx0:RH7:D x0complementary F x /\ D x0X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F xexists z : R, D zX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x1 : R, F x1 -> exists y : R, f0 y x1H4:family_open_set f0x:RH5:X xH6:~ F xx0:RH7:D x0complementary F xX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x1 : R, F x1 -> exists y : R, f0 y x1H4:family_open_set f0x:RH5:X xH6:~ F xx0:RH7:D x0D x0X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F xexists z : R, D zX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x1 : R, F x1 -> exists y : R, f0 y x1H4:family_open_set f0x:RH5:X xH6:~ F xx0:RH7:D x0D x0X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F xexists z : R, D zX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F xexists z : R, D zX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F xz0:RH7:F z0exists z : R, D zX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:forall x0 : R, F x0 -> exists y : R, f0 y x0H4:family_open_set f0x:RH5:X xH6:~ F xz0:RH7:F z0H8:exists y : R, f0 y z0exists z : R, D zX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> PropH3:forall x : R, (exists y : R, g' x y) -> D' xf':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0family_open_set f'X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xopen_set (union_domain (f0 x) (complementary F))X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xunion_domain (f0 x) (complementary F) =_D (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xopen_set (f0 x)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xopen_set (complementary F)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xunion_domain (f0 x) (complementary F) =_D (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xopen_set (complementary F)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xunion_domain (f0 x) (complementary F) =_D (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xunion_domain (f0 x) (complementary F) =_D (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xincluded (union_domain (f0 x) (complementary F)) (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xincluded (fun y : R => f0 x y \/ complementary F y /\ D x) (union_domain (f0 x) (complementary F))X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xx0:RH6:f0 x x0 \/ ~ F x0f0 x x0 \/ ~ F x0 /\ D xX, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xincluded (fun y : R => f0 x y \/ complementary F y /\ D x) (union_domain (f0 x) (complementary F))X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xincluded (fun y : R => f0 x y \/ complementary F y /\ D x) (union_domain (f0 x) (complementary F))X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:D xx0:RH6:f0 x x0 \/ ~ F x0 /\ D xf0 x x0 \/ ~ F x0X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xopen_set (f0 x)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xf0 x =_D (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xf0 x =_D (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xincluded (f0 x) (fun y : R => f0 x y \/ complementary F y /\ D x)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xincluded (fun y : R => f0 x y \/ complementary F y /\ D x) (f0 x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> PropH3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xincluded (fun y : R => f0 x y \/ complementary F y /\ D x) (f0 x)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xx0:RH6:f0 x x0 \/ ~ F x0 /\ D xf0 x x0X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xx0:RH6:f0 x x0 \/ ~ F x0 /\ D xH7:f0 x x0f0 x x0X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xx0:RH6:f0 x x0 \/ ~ F x0 /\ D xH7:~ F x0 /\ D xf0 x x0X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyHyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> PropD':=D:R -> PropH3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1f':={| ind := D'; f := g'; cond_fam := H3 |}:familyH2:covering F f0H4:family_open_set f0x:RH5:~ D xx0:RH6:f0 x x0 \/ ~ F x0 /\ D xH7:~ F x0 /\ D xf0 x x0X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> PropD':=D:R -> Propforall x : R, (exists y : R, g' x y) -> D' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> Propx:RH3:exists y : R, g' x yy0:RH4:f0 x y0 \/ complementary F y0 /\ D xH5:f0 x y0D' xX, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> Propx:RH3:exists y : R, g' x yy0:RH4:f0 x y0 \/ complementary F y0 /\ D xH5:complementary F y0 /\ D xD' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH: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 Xf0:familyH2:covering_open_set F f0Hyp_F_NE:exists z : R, F zD:=ind f0:R -> Propg:=f f0:R -> R -> Propg':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> PropD':=D:R -> Propx:RH3:exists y : R, g' x yy0:RH4:f0 x y0 \/ complementary F y0 /\ D xH5:complementary F y0 /\ D xD' xX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_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 -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)compact F -> exists D : R -> Prop, covering_finite F (subfamily f0 D)X, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)compact FX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)compact FX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)compact (fun _ : R => False)X, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)(fun _ : R => False) =_D FX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)(fun _ : R => False) =_D FX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)included (fun _ : R => False) FX, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_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. (**********)X, F:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)H0:closed_set FH1:included F Xf0:familyH2:covering_open_set F f0Hyp_F_NE:~ (exists z : R, F z)included F (fun _ : R => False)forall X : R -> Prop, closed_set X -> bounded X -> compact Xforall X : R -> Prop, closed_set X -> bounded X -> compact XX:R -> PropH:closed_set XH0:exists m M : R, forall x : R, X x -> m <= x <= Mcompact XX:R -> PropH:closed_set Xm:RH0:exists M : R, forall x : R, X x -> m <= x <= Mcompact XX:R -> PropH:closed_set Xm, M:RH0:forall x : R, X x -> m <= x <= Mcompact Xapply (compact_P4 (fun c:R => m <= c <= M) X H1 H H0). Qed. (**********)X:R -> PropH:closed_set Xm, M:RH0:forall x : R, X x -> m <= x <= MH1:compact (fun c : R => m <= c <= M)compact Xforall X : R -> Prop, compact X <-> closed_set X /\ bounded Xforall X : R -> Prop, compact X <-> closed_set X /\ bounded XX:R -> Propcompact X -> closed_set X /\ bounded XX:R -> Propclosed_set X /\ bounded X -> compact Xintro; 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. (**********)X:R -> Propclosed_set X /\ bounded X -> compact Xforall (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 -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D : R -> Prop, covering_finite X (subfamily f2 D)f1:familyH1:covering (image_dir f0 X) f1 /\ family_open_set f1exists D : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D)f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D : R -> Prop, covering_finite X (subfamily f2 D)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1exists D : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D)f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propexists D0 : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D0)f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propexists D0 : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D0)f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=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 -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familyexists D0 : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D0)f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familycovering_open_set X f' -> exists D0 : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D0)f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familycovering_open_set X f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f'D':R -> PropH5:covering_finite X (subfamily f' D')covering_finite (image_dir f0 X) (subfamily f1 D')f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familycovering_open_set X f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f'D':R -> PropH5:covering X (subfamily f' D')H6:family_finite (subfamily f' D')covering (image_dir f0 X) (subfamily f1 D')f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f'D':R -> PropH5:covering X (subfamily f' D')H6:family_finite (subfamily f' D')family_finite (subfamily f1 D')f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familycovering_open_set X f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f'D':R -> PropH5:covering X (subfamily f' D')H6:family_finite (subfamily f' D')family_finite (subfamily f1 D')f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familycovering_open_set X f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x0 : R, continuity_pt f0 x0H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f':={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f'D':R -> PropH5:covering X (subfamily f' D')H6:exists l0 : Rlist, forall x0 : R, ind (subfamily f' D') x0 <-> In x0 l0l:RlistH7:forall x0 : R, ind (subfamily f' D') x0 <-> In x0 lx:RH8:ind (subfamily f' D') x -> In x lH9:In x l -> ind (subfamily f' D') xH10:ind (subfamily f1 D') xIn x lf0:R -> RX:R -> PropH:forall x0 : R, continuity_pt f0 x0H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f':={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f'D':R -> PropH5:covering X (subfamily f' D')H6:exists l0 : Rlist, forall x0 : R, ind (subfamily f' D') x0 <-> In x0 l0l:RlistH7:forall x0 : R, ind (subfamily f' D') x0 <-> In x0 lx:RH8:ind (subfamily f' D') x -> In x lH9:In x l -> ind (subfamily f' D') xH10:In x lind (subfamily f1 D') xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familycovering_open_set X f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x0 : R, continuity_pt f0 x0H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f':={| ind := D; f := g; cond_fam := H3 |}:familyH4:covering_open_set X f'D':R -> PropH5:covering X (subfamily f' D')H6:exists l0 : Rlist, forall x0 : R, ind (subfamily f' D') x0 <-> In x0 l0l:RlistH7:forall x0 : R, ind (subfamily f' D') x0 <-> In x0 lx:RH8:ind (subfamily f' D') x -> In x lH9:In x l -> ind (subfamily f' D') xH10:In x lind (subfamily f1 D') xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familycovering_open_set X f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familycovering_open_set X f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familycovering X f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x0 : R, continuity_pt f0 x0H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:forall x0 : R, (exists y : R, x0 = f0 y /\ X y) -> exists y : R, f1 y x0H2:family_open_set f1D:=ind f1:R -> Propg:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f':={| ind := D; f := g; cond_fam := H3 |}:familyx:RH4:X xexists y : R, f0 x = f0 y /\ X yf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> PropH3:forall x : R, (exists y : R, g x y) -> D xf':={| ind := D; f := g; cond_fam := H3 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x0 : R, continuity_pt f0 x0H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:forall x0 : R, open_set (f1 x0)D:=ind f1:R -> Propg:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f':={| ind := D; f := g; cond_fam := H3 |}:familyx: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 -> RX:R -> PropH:forall x0 : R, continuity_pt f0 x0H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:forall x0 : R, open_set (f1 x0)D:=ind f1:R -> Propg:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f':={| ind := D; f := g; cond_fam := H3 |}:familyx:R(fun y : R => image_rec f0 (f1 x) y) = image_rec f0 (f1 x)f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x0 : R, continuity_pt f0 x0H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:forall x0 : R, open_set (f1 x0)D:=ind f1:R -> Propg:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f':={| ind := D; f := g; cond_fam := H3 |}:familyx:RH4:(fun y : R => image_rec f0 (f1 x) y) = image_rec f0 (f1 x)open_set (image_rec f0 (f1 x))f0:R -> RX:R -> PropH:forall x0 : R, continuity_pt f0 x0H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:forall x0 : R, open_set (f1 x0)D:=ind f1:R -> Propg:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f':={| ind := D; f := g; cond_fam := H3 |}:familyx:R(fun y : R => image_rec f0 (f1 x) y) = image_rec f0 (f1 x)f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xf0:R -> RX:R -> PropH:forall x0 : R, continuity_pt f0 x0H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:forall x0 : R, open_set (f1 x0)D:=ind f1:R -> Propg:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> PropH3:forall x0 : R, (exists y : R, g x0 y) -> D x0f':={| ind := D; f := g; cond_fam := H3 |}:familyx:R(fun y : R => image_rec f0 (f1 x) y) = image_rec f0 (f1 x)f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xintros; apply (cond_fam f1); unfold g in H3; unfold image_rec in H3; elim H3; intros; exists (f0 x0); apply H4. Qed.f0:R -> RX:R -> PropH:forall x : R, continuity_pt f0 xH0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)f1:familyH1:covering (image_dir f0 X) f1H2:family_open_set f1D:=ind f1:R -> Propg:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Propforall x : R, (exists y : R, g x y) -> D xforall (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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> Rexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R0 < b - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - acontinuity hf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x < acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0a - x > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < a - x -> Rabs (h x0 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < a - x -> Rabs (h x0 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xRabs ((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)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xr:x <= aRabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xr:x <= ar0:x0 <= aRabs (f0 a - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xr:x <= ax0 <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xr:x <= ax0 <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xr:x <= ax0 + - x <= Rabs (x0 - x)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xr:x <= aRabs (x0 - x) < a + - xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xr:x <= aRabs (x0 - x) < a + - xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x < aeps:RH4:eps > 0x0:RH5:Rabs (x0 - x) < a - xx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aa <= a <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8: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) > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8: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) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)r:x0 <= b - ax0 > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)n:~ x0 <= b - ab - a > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8: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) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)n:~ x0 <= b - ab - a > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8: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) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8: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) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < b -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aRabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHlta:x1 <= aRabs (f0 a - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bRabs (f0 x1 - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsD_x no_cond a x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsRabs (x1 - a) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsTruef0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsa <> x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsRabs (x1 - a) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsa <> x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsRabs (x1 - a) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsRabs (x1 - a) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsRabs (x1 - a) < Rmin x0 (b - a)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsRmin x0 (b - a) <= x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= aHleb:x1 <= bH11:x0 > 0H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < epsRmin x0 (b - a) <= x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < br:x <= aHnlea:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)H10:x1 < bx <= af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - x) < Rmin x0 (b - a)x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - a) < Rmin x0 (b - a)x1 + - a <= Rabs (x1 - a)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - a) < Rmin x0 (b - a)Rabs (x1 - a) < b + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - a) < Rmin x0 (b - a)Rabs (x1 - a) < b + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - a) < Rmin x0 (b - a)Rabs (x1 - a) < Rmin x0 (b - a)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - a) < Rmin x0 (b - a)Rmin x0 (b - a) <= b + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x = aH5:a <= a <= bH6: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:RH7:eps > 0x0:RH8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)x1:RH9:Rabs (x1 - a) < Rmin x0 (b - a)Rmin x0 (b - a) <= b + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > acontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < ba <= x <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsexists 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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps0 < x - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aexists 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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aexists 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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - a0 < b - xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xexists 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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xexists 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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xRmin x0 (Rmin (x - a) (b - x)) > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xforall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHle:x - a <= b - x(if Rle_dec x0 (x - a) then x0 else x - a) > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHnle:~ x - a <= b - x(if Rle_dec x0 (b - x) then x0 else b - x) > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xforall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHle:x - a <= b - xHlea:x0 <= x - ax0 > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHle:x - a <= b - xHnlea:~ x0 <= x - ax - a > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHnle:~ x - a <= b - x(if Rle_dec x0 (b - x) then x0 else b - x) > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xforall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHle:x - a <= b - xHnlea:~ x0 <= x - ax - a > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHnle:~ x - a <= b - x(if Rle_dec x0 (b - x) then x0 else b - x) > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xforall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHnle:~ x - a <= b - x(if Rle_dec x0 (b - x) then x0 else b - x) > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xforall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHnle:~ x - a <= b - xHleb:x0 <= b - xx0 > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHnle:~ x - a <= b - xHnleb:~ x0 <= b - xb - x > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xforall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xHnle:~ x - a <= b - xHnleb:~ x0 <= b - xb - x > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xforall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < epsH11:0 < x - aH12:0 < b - xforall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < b -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHle:x <= aRabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= aRabs ((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)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= aRabs ((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)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bRabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bHle0:x1 <= aRabs (f0 a - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= aRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= aRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ar0:x1 <= bRabs (f0 x1 - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ar0:x1 <= bD_x no_cond x x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ar0:x1 <= bRabs (x1 - x) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ar0:x1 <= bRabs (x1 - x) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ar0:x1 <= bRabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ar0:x1 <= bRmin x0 (Rmin (x - a) (b - x)) <= x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ar0:x1 <= bRmin x0 (Rmin (x - a) (b - x)) <= x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ar:x <= bn:~ x1 <= ax1 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))H15:a < x1H16:x1 < bHnle:~ x <= ax <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x + - x1 <= Rabs (x1 - x)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rabs (x1 - x) < x + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rabs (x1 - x) < x + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rmin x0 (Rmin (x - a) (b - x)) <= x + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rmin x0 (Rmin (x - a) (b - x)) <= x + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rmin x0 (Rmin (x - a) (b - x)) <= Rmin (x - a) (b - x)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rmin (x - a) (b - x) <= x + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rmin (x - a) (b - x) <= x + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x1 < bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))x1 + - x <= Rabs (x1 - x)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rabs (x1 - x) < b + - xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rabs (x1 - x) < b + - xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rmin x0 (Rmin (x - a) (b - x)) <= b + - xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x < bH6:a <= x <= bH7: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:RH8:eps > 0x0:RH9:x0 > 0H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < epsH11:0 < x - aH12:0 < b - xx1:RH13:D_x no_cond x x1H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))Rmin x0 (Rmin (x - a) (b - x)) <= b + - xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = ba <= b <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10: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) > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10: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) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)r:x0 <= b - ax0 > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)n:~ x0 <= b - ab - a > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10: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) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)n:~ x0 <= b - ab - a > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10: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) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10: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) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1 -> Rabs (h x1 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hlea:x <= aRabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aRabs ((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)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aRabs ((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)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHlea':x1 <= aRabs (f0 a - (if Rle_dec x b then f0 x else f0 b)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHleb':x1 <= bRabs (f0 x1 - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (f0 x1 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 = bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (f0 x1 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsD_x no_cond b x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (x1 - b) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 = bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (f0 x1 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsTruef0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsb <> x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (x1 - b) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 = bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (f0 x1 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsb <> x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (x1 - b) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 = bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (f0 x1 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (x1 - b) < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 = bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (f0 x1 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - b) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (x1 - b) < Rmin x0 (b - a)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - b) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRmin x0 (b - a) <= x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 = bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (f0 x1 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - b) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 < bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRmin x0 (b - a) <= x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 = bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (f0 x1 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bH15:x1 = bH13:x0 > 0H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < epsRabs (f0 x1 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHleb:x <= bHnleb':~ x1 <= bRabs (f0 b - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)H12:a < x1Hnlea:~ x <= aHnlea':~ x1 <= aHnleb:~ x <= bRabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - x) < Rmin x0 (b - a)a < x1f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - b) < Rmin x0 (b - a)b + - x1 <= Rabs (x1 - b)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - b) < Rmin x0 (b - a)Rabs (x1 - b) < b + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - b) < Rmin x0 (b - a)Rabs (x1 - b) < b + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - b) < Rmin x0 (b - a)Rabs (x1 - b) < Rmin x0 (b - a)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - b) < Rmin x0 (b - a)Rmin x0 (b - a) <= b + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x = bH7:a <= b <= bH8: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:RH9:eps > 0x0:RH10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)x1:RH11:Rabs (x1 - b) < Rmin x0 (b - a)Rmin x0 (b - a) <= b + - af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > bcontinuity_pt h xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x - b > 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < x - b -> Rabs (h x0 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < x - b -> Rabs (h x0 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bRabs (h x0 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bb < x0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Rabs (h x0 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bx + - x0 <= Rabs (x0 - x)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bRabs (x0 - x) < x + - bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Rabs (h x0 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bRabs (x0 - x) < x + - bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Rabs (h x0 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Rabs (h x0 - h x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hle:x <= aRabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 a) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aRabs ((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)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aRabs ((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)) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aHleb:x <= bRabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 x) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aHnleb:~ x <= bRabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aHnleb:~ x <= bRabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aHnleb:~ x <= bHlea':x0 <= aRabs (f0 a - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aHnleb:~ x <= bHnlea':~ x0 <= aRabs ((if Rle_dec x0 b then f0 x0 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aHnleb:~ x <= bHnlea':~ x0 <= aRabs ((if Rle_dec x0 b then f0 x0 else f0 b) - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aHnleb:~ x <= bHnlea':~ x0 <= aHleb':x0 <= bRabs (f0 x0 - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aHnleb:~ x <= bHnlea':~ x0 <= aHnleb':~ x0 <= bRabs (f0 b - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> RH2:0 < b - ax:RH3:x = a \/ x > aH4:x > aH5:x = b \/ x > bH6:x > beps:RH7:eps > 0x0:RH8:D_x no_cond x x0H9:Rabs (x0 - x) < x - bH10:b < x0Hnle:~ x <= aHnleb:~ x <= bHnlea':~ x0 <= aHnleb':~ x0 <= bRabs (f0 b - f0 b) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - aforall c : R, a <= c <= b -> h c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - ac:RH3:a <= c <= bH4:a <= cH5:c <= bH6:c < af0 a = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - ac:RH3:a <= c <= bH4:a <= cH5:c <= bH6:c = af0 a = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - ac:RH3:a <= c <= bH4:a <= cH5:c <= bn:~ c <= a(if Rle_dec c b then f0 c else f0 b) = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - ac:RH3:a <= c <= bH4:a <= cH5:c <= bH6:c = af0 a = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - ac:RH3:a <= c <= bH4:a <= cH5:c <= bn:~ c <= a(if Rle_dec c b then f0 c else f0 b) = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - ac:RH3:a <= c <= bH4:a <= cH5:c <= bn:~ c <= a(if Rle_dec c b then f0 c else f0 b) = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - ac:RH3:a <= c <= bH4:a <= cH5:c <= bn:~ c <= ar:c <= bf0 c = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - ac:RH3:a <= c <= bH4:a <= cH5:c <= bn:~ c <= ac <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a < bh:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> RH2:0 < b - ac:RH3:a <= c <= bH4:a <= cH5:c <= bn:~ c <= ac <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bcontinuity (fun _ : R => f0 a)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bforall c : R, a <= c <= b -> f0 a = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:a = bforall c : R, a <= c <= b -> f0 a = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a = bc:RH2:a <= c <= bH3:b <= cH4:c <= bb = c -> f0 a = f0 cf0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a = bc:RH2:a <= c <= bH3:b <= cH4:c <= bb = capply Rle_antisym; assumption. Qed. (**********)f0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0H1:a = bc:RH2:a <= c <= bH3:b <= cH4:c <= bb = cforall (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 <= bforall (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 <= bf0:R -> Ra, b:RH:a <= bH0: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont_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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cexists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1:compact (fun c : R => a <= c <= b)exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) Mimage_dir g (fun c : R => a <= c <= b) M -> exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) Mimage_dir g (fun c : R => a <= c <= b) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MMxx:RH8:M = g MxxH9:a <= Mxx <= bforall c : R, a <= c <= b -> f0 c <= f0 Mxxf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MMxx:RH8:M = g MxxH9:a <= Mxx <= ba <= Mxx <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) Mimage_dir g (fun c : R => a <= c <= b) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MMxx:RH8:M = g MxxH9:a <= Mxx <= ba <= Mxx <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) Mimage_dir g (fun c : R => a <= c <= b) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) Mimage_dir g (fun c : R => a <= c <= b) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:image_dir g (fun c : R => a <= c <= b) Mimage_dir g (fun c : R => a <= c <= b) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mimage_dir g (fun c : R => a <= c <= b) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mimage_dir g (fun c : R => a <= c <= b) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ 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) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps) -> image_dir g (fun c : R => a <= c <= b) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0H11:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)H12:M <= M - epsM - eps < M -> image_dir g (fun c : R => a <= c <= b) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0H11:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)H12:M <= M - epsM - eps < Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0H11:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)H12:M <= M - epsM - eps < Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= M -> x <= M - epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= MH13:x <= M - epsx <= M - epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsx <= M - epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsx <= M - epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsRabs (x - M) < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsexists y : R, x = g y /\ a <= y <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - x < epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - x >= 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsexists y : R, x = g y /\ a <= y <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - eps < x - eps + epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - eps = x - eps + (M - x)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - x >= 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsexists y : R, x = g y /\ a <= y <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - eps < xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsx = x - eps + epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - eps = x - eps + (M - x)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - x >= 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsexists y : R, x = g y /\ a <= y <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsx = x - eps + epsf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - eps = x - eps + (M - x)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - x >= 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsexists y : R, x = g y /\ a <= y <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - eps = x - eps + (M - x)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - x >= 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsexists y : R, x = g y /\ a <= y <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsM - x >= 0f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsexists y : R, x = g y /\ a <= y <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bH12:x <= Mn:~ x <= M - epsexists y : R, x = g y /\ a <= y <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) x0H6:bound (image_dir g (fun c : R => a <= c <= b))M:RH8:~ image_dir g (fun c : R => a <= c <= b) Meps:posrealH9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yH7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) MH10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0x:RH11:exists y : R, x = g y /\ a <= y <= bx <= Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ 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)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9: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 -> PropH10:neighbourhood V MH11:forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) yexists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9: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 -> PropH10:exists delta : posreal, included (disc M delta) VH11:forall y0 : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y0del:posrealH12:included (disc M del) Vy:RH13:intersection_domain (disc M del) (image_dir g (fun c : R => a <= c <= b)) yintersection_domain V (image_dir g (fun c : R => a <= c <= b)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9: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 -> PropH10:exists delta : posreal, included (disc M delta) VH11:forall y0 : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y0del:posrealH12:included (disc M del) Vy:RH13:disc M del yH14:image_dir g (fun c : R => a <= c <= b) yV yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9: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 -> PropH10:exists delta : posreal, included (disc M delta) VH11:forall y0 : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y0del:posrealH12:included (disc M del) Vy:RH13:disc M del yH14:image_dir g (fun c : R => a <= c <= b) yimage_dir g (fun c : R => a <= c <= b) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9: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 -> PropH10:exists delta : posreal, included (disc M delta) VH11:forall y0 : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y0del:posrealH12:included (disc M del) Vy:RH13:disc M del yH14:image_dir g (fun c : R => a <= c <= b) yimage_dir g (fun c : R => a <= c <= b) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) Mexists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ 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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) M~ point_adherent (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:~ (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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) M~ point_adherent (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:~ (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) nexists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) M~ point_adherent (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:~ (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) nV0:R -> PropH11:~ (neighbourhood V0 M -> exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)H12:neighbourhood V0 MH13:~ (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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) M~ point_adherent (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:~ (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) nV0:R -> PropH11:~ (neighbourhood V0 M -> exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)H12:neighbourhood V0 MH13:~ (exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)neighbourhood V0 Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:~ (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) nV0:R -> PropH11:~ (neighbourhood V0 M -> exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)H12:neighbourhood V0 MH13:~ (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)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) M~ point_adherent (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:~ (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) nV0:R -> PropH11:~ (neighbourhood V0 M -> exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)H12:neighbourhood V0 MH13:~ (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)) yf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) M~ point_adherent (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) M~ point_adherent (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:point_adherent (image_dir g (fun c : R => a <= c <= b)) Madherence (image_dir g (fun c : R => a <= c <= b)) M -> Falsef0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:point_adherent (image_dir g (fun c : R => a <= c <= b)) Madherence (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:point_adherent (image_dir g (fun c : R => a <= c <= b)) MH10:adherence (image_dir g (fun c : R => a <= c <= b)) MH11: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))Falsef0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:point_adherent (image_dir g (fun c : R => a <= c <= b)) Madherence (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:point_adherent (image_dir g (fun c : R => a <= c <= b)) MH10:adherence (image_dir g (fun c : R => a <= c <= b)) MH11: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) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:point_adherent (image_dir g (fun c : R => a <= c <= b)) Madherence (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:point_adherent (image_dir g (fun c : R => a <= c <= b)) MH10:adherence (image_dir g (fun c : R => a <= c <= b)) MH11: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) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:point_adherent (image_dir g (fun c : R => a <= c <= b)) Madherence (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xH6:bound (image_dir g (fun c : R => a <= c <= b))M:RH7:is_lub (image_dir g (fun c : R => a <= c <= b)) MH8:~ image_dir g (fun c : R => a <= c <= b) MH9:point_adherent (image_dir g (fun c : R => a <= c <= b)) Madherence (image_dir g (fun c : R => a <= c <= b)) Mf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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) xf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 af0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cHypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)g:R -> RHcont:continuity gHeq:forall c : R, a <= c <= b -> g c = f0 cH1: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 -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)apply prolongement_C0; assumption. Qed. (**********)f0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists 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 mx : R, (forall c : R, a <= c <= b -> f0 mx <= f0 c) /\ a <= mx <= bforall (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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cexists mx : R, (forall c : R, a <= c <= b -> f0 mx <= f0 c) /\ a <= mx <= bf0:R -> Ra, b:RH:a <= bH0: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 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cforall c : R, a <= c <= b -> continuity_pt (- f0) cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:forall c : R, a <= c <= b -> continuity_pt (- f0) cH2:exists Mx : R, (forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F Mx) /\ a <= Mx <= bx0:RH3:(forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F x0) /\ a <= x0 <= bforall c : R, a <= c <= b -> f0 x0 <= f0 cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:forall c : R, a <= c <= b -> continuity_pt (- f0) cH2:exists Mx : R, (forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F Mx) /\ a <= Mx <= bx0:RH3:(forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F x0) /\ a <= x0 <= ba <= x0 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cforall c : R, a <= c <= b -> continuity_pt (- f0) cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cH1:forall c : R, a <= c <= b -> continuity_pt (- f0) cH2:exists Mx : R, (forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F Mx) /\ a <= Mx <= bx0:RH3:(forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F x0) /\ a <= x0 <= ba <= x0 <= bf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cforall c : R, a <= c <= b -> continuity_pt (- f0) cf0:R -> Ra, b:RH:a <= bH0:forall c : R, a <= c <= b -> continuity_pt f0 cforall c : R, a <= c <= b -> continuity_pt (- f0) cf0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0c:RH1:a <= c <= bcontinuity_pt (- f0) capply (continuity_pt_opp _ _ H2). Qed. (********************************************************)f0:R -> Ra, b:RH:a <= bH0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0c:RH1:a <= c <= bH2:continuity_pt f0 ccontinuity_pt (- f0) c
(********************************************************) 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 xforall 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 xun:nat -> RD:=fun x1 : R => exists n : nat, x1 = INR n:R -> Propf0:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Propx:RH:exists y : R, f0 x yx0:RH0:f0 x x0neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0un:nat -> RD:=fun x1 : R => exists n : nat, x1 = INR n:R -> Propf0:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Propx:RH:exists y : R, f0 x yx0:RH0:f0 x x0H1:neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0D xelim (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)).un:nat -> RD:=fun x1 : R => exists n : nat, x1 = INR n:R -> Propf0:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Propx:RH:exists y : R, f0 x yx0:RH0:f0 x x0H1:neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0D xforall (un : nat -> R) (x : R), ValAdh un x <-> ValAdh_un un xforall (un : nat -> R) (x : R), ValAdh un x <-> ValAdh_un un xun:nat -> Rx:RH:ValAdh un xValAdh_un un xun:nat -> Rx:RH:ValAdh_un un xValAdh un xun:nat -> Rx:RH:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)y:RH0:exists n : nat, y = INR nN:natH1:y = INR NV:R -> PropH2:neighbourhood V xx0:natH3:(N <= x0)%natH4:V (un x0)V (un x0)un:nat -> Rx:RH:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)y:RH0:exists n : nat, y = INR nN:natH1:y = INR NV:R -> PropH2:neighbourhood V xx0:natH3:(N <= x0)%natH4:V (un x0)(exists p : nat, un x0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)un:nat -> Rx:RH:ValAdh_un un xValAdh un xun:nat -> Rx:RH:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)y:RH0:exists n : nat, y = INR nN:natH1:y = INR NV:R -> PropH2:neighbourhood V xx0:natH3:(N <= x0)%natH4:V (un x0)(exists p : nat, un x0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)un:nat -> Rx:RH:ValAdh_un un xValAdh un xun:nat -> Rx:RH:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)y:RH0:exists n : nat, y = INR nN:natH1:y = INR NV:R -> PropH2:neighbourhood V xx0:natH3:(N <= x0)%natH4:V (un x0)exists p : nat, un x0 = un p /\ y <= INR pun:nat -> Rx:RH:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)y:RH0:exists n : nat, y = INR nN:natH1:y = INR NV:R -> PropH2:neighbourhood V xx0:natH3:(N <= x0)%natH4:V (un x0)exists n : nat, y = INR nun:nat -> Rx:RH:ValAdh_un un xValAdh un xun:nat -> Rx:RH:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)y:RH0:exists n : nat, y = INR nN:natH1:y = INR NV:R -> PropH2:neighbourhood V xx0:natH3:(N <= x0)%natH4:V (un x0)exists n : nat, y = INR nun:nat -> Rx:RH:ValAdh_un un xValAdh un xun:nat -> Rx:RH:ValAdh_un un xValAdh un xun:nat -> Rx:RH: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)) xV:R -> PropN:natH0:neighbourhood V xadherence (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) xun:nat -> Rx:RH: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)) xV:R -> PropN:natH0:neighbourhood V xH1:adherence (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) xexists p : nat, (N <= p)%nat /\ V (un p)un:nat -> Rx:RH: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)) xV:R -> PropN:natH0:neighbourhood V xH1:adherence (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) xexists p : nat, (N <= p)%nat /\ V (un p)un:nat -> Rx:RH: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)) xV:R -> PropN:natH0:neighbourhood V xH1: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)) yH2: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)) yx0:RH3:V x0H5:exists n : nat, INR N = INR nx1:natH4:x0 = un x1H6:INR N <= INR x1(N <= x1)%natun:nat -> Rx:RH: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)) xV:R -> PropN:natH0:neighbourhood V xH1: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)) yH2: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)) yx0:RH3:V x0H5:exists n : nat, INR N = INR nx1:natH4:x0 = un x1H6:INR N <= INR x1V (un x1)rewrite H4 in H3; apply H3. Qed.un:nat -> Rx:RH: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)) xV:R -> PropN:natH0:neighbourhood V xH1: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)) yH2: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)) yx0:RH3:V x0H5:exists n : nat, INR N = INR nx1:natH4:x0 = un x1H6:INR N <= INR x1V (un x1)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 F G : R -> Prop, included F G -> included (adherence F) (adherence G)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 -> PropH:compact XHyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gexists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:compact XHyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propexists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:compact XHyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Propexists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:compact XHyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Propforall x : R, (exists y : R, f' x y) -> D' xX:R -> PropH:compact XHyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xexists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:compact XHyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xexists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:compact XHyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyexists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familycovering_open_set X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familycovering X f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyfamily_open_set f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyfamily_open_set f0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:D' xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:~ D' xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:D' xopen_set (complementary (g x))X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:D' xcomplementary (g x) =_D f0 xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:~ D' xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:D' xcomplementary (g x) =_D f0 xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:~ D' xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:D' xincluded (complementary (g x)) (fun y : R => complementary (g x) y /\ D' x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:D' xincluded (fun y : R => complementary (g x) y /\ D' x) (complementary (g x))X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:~ D' xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:D' xincluded (fun y : R => complementary (g x) y /\ D' x) (complementary (g x))X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:~ D' xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:~ D' xopen_set (f0 x)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:~ D' xopen_set (fun _ : R => False)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:~ D' x(fun _ : R => False) =_D f0 xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyx:RH3:~ D' x(fun _ : R => False) =_D f0 xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)intersection_vide_in X (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:Rintersection_domain (ind g) SF x -> included (fun y : R => g x y /\ SF x) XX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:R~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x1 : R, (ind g x1 -> included (g x1) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RH5:intersection_domain (ind g) SF xx0:RH6:g x x0 /\ SF xH7:ind g x -> included (g x) XH8:~ (exists y : R, intersection_family g y)H9:g x x0H10:SF xind g xX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x1 : R, (ind g x1 -> included (g x1) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RH5:intersection_domain (ind g) SF xx0:RH6:g x x0 /\ SF xH7:ind g x -> included (g x) XH8:~ (exists y : R, intersection_family g y)H9:g x x0H10:SF xg x x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:R~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x1 : R, (ind g x1 -> included (g x1) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RH5:intersection_domain (ind g) SF xx0:RH6:g x x0 /\ SF xH7:ind g x -> included (g x) XH8:~ (exists y : R, intersection_family g y)H9:g x x0H10:SF xg x x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:R~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:R~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF y~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yFalseX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yX x0 -> FalseX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yX x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF yH7:X x0H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x1:RH8:f' x1 x0 /\ SF x1ind g x1 /\ SF x1 -> FalseX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF yH7:X x0H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x1:RH8:f' x1 x0 /\ SF x1ind g x1 /\ SF x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yX x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF yH7:X x0H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x1:RH8:f' x1 x0 /\ SF x1ind g x1 /\ SF x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yX x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF yH7:X x0H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x1:RH8:f' x1 x0 /\ SF x1ind g x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF yH7:X x0H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x1:RH8:f' x1 x0 /\ SF x1SF x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yX x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF yH7:X x0H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x1:RH8:f' x1 x0 /\ SF x1exists y : R, f0 x1 yX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF yH7:X x0H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x1:RH8:f' x1 x0 /\ SF x1SF x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yX x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF yH7:X x0H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x1:RH8:f' x1 x0 /\ SF x1SF x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yX x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> PropH2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yX x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yx1:RH7:intersection_domain (ind g) SF x1H8:g x1 x0 /\ SF x1H9:g x1 x0H10:SF x1ind g x1 -> X x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yx1:RH7:intersection_domain (ind g) SF x1H8:g x1 x0 /\ SF x1H9:g x1 x0H10:SF x1ind g x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yx1:RH7:intersection_domain (ind g) SF x1H8:g x1 x0 /\ SF x1H9:g x1 x0H10:SF x1H11:ind g x1H12:ind g x1 -> included (g x1) XH13:~ (exists y : R, intersection_family g y)ind g x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yx1:RH7:intersection_domain (ind g) SF x1H8:g x1 x0 /\ SF x1H9:g x1 x0H10:SF x1H11:ind g x1H12:ind g x1 -> included (g x1) XH13:~ (exists y : R, intersection_family g y)g x1 x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yx1:RH7:intersection_domain (ind g) SF x1H8:g x1 x0 /\ SF x1H9:g x1 x0H10:SF x1ind g x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yx1:RH7:intersection_domain (ind g) SF x1H8:g x1 x0 /\ SF x1H9:g x1 x0H10:SF x1H11:ind g x1H12:ind g x1 -> included (g x1) XH13:~ (exists y : R, intersection_family g y)g x1 x0X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yx1:RH7:intersection_domain (ind g) SF x1H8:g x1 x0 /\ SF x1H9:g x1 x0H10:SF x1ind g x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)D':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':exists y : R, intersection_domain (ind g) SF yH5:exists y : R, intersection_family (subfamily g SF) yx0:RH6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF yx1:RH7:intersection_domain (ind g) SF x1H8:g x1 x0 /\ SF x1H9:g x1 x0H10:SF x1ind g x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)x:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)~ (exists y : R, intersection_family (subfamily g SF) y)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':~ (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 -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)H4:covering X (subfamily f0 SF)exists z : R, X zX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x0:RH5:X x0x1:RH6:f' x1 x0 /\ SF x1H7:f' x1 x0H8:SF x1ind g x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x0:RH5:X x0x1:RH6:f' x1 x0 /\ SF x1H7:f' x1 x0H8:SF x1SF x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)H4:covering X (subfamily f0 SF)exists z : R, X zX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> PropH2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2x0:RH5:X x0x1:RH6:f' x1 x0 /\ SF x1H7:f' x1 x0H8:SF x1SF x1X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)H4:covering X (subfamily f0 SF)exists z : R, X zX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> PropH2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0f0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> Propx:RHyp':~ (exists y : R, intersection_domain (ind g) SF y)H4:covering X (subfamily f0 SF)exists z : R, X zX:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4: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.X:R -> PropH:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)Hyp:exists z : R, X zg:familyH0:family_closed_set gH1:intersection_vide_in X gD':=ind g:R -> Propf':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> PropH2:forall x : R, (exists y : R, f' x y) -> D' xf0:={| ind := D'; f := f'; cond_fam := H2 |}:familyH3:covering_open_set X f0SF:R -> PropH4:covering_finite X (subfamily f0 SF)family_finite (subfamily g SF)forall (un : nat -> R) (X : R -> Prop), compact X -> (forall n : nat, X (un n)) -> exists l : R, ValAdh un lforall (un : nat -> R) (X : R -> Prop), compact X -> (forall n : nat, X (un n)) -> exists l : R, ValAdh un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)(exists l : R, ValAdh_un un l) -> exists l : R, ValAdh un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)exists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)exists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)exists z : R, X zun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Propexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Propforall x : R, (exists y : R, g x y) -> D xun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x1 : R => exists n : nat, x1 = INR n:R -> Propg:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Propx:RH2:exists y : R, g x yx0:RH3: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) yD xun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x1 : R => exists n : nat, x1 = INR n:R -> Propg:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Propx:RH2:exists y : R, g x yx0:RH3: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) yneighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x1 : R => exists n : nat, x1 = INR n:R -> Propg:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Propx:RH2:exists y : R, g x yx0:RH3: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) yH4:neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0D xun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x1 : R => exists n : nat, x1 = INR n:R -> Propg:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Propx:RH2:exists y : R, g x yx0:RH3: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) yH4:neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0D xun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 lexists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0 -> exists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)exists l : R, ValAdh_un un lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RD r -> exists y : R, intersection_family (subfamily f0 SF) yun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RD run:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 lr:=MaxRlist l:RH11:exists n : nat, r = INR nx:natH12:r = INR xy:RH13:D y /\ SF yg y (un x)un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 lr:=MaxRlist l:RH11:exists n : nat, r = INR nx:natH12:r = INR xy:RH13:D y /\ SF ySF yun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RD run:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 lr:=MaxRlist l:RH11:exists n : nat, r = INR nx:natH12:r = INR xy:RH13:D y /\ SF yexists p : nat, un x = un p /\ y <= INR pun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 lr:=MaxRlist l:RH11:exists n : nat, r = INR nx:natH12:r = INR xy:RH13:D y /\ SF yD yun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 lr:=MaxRlist l:RH11:exists n : nat, r = INR nx:natH12:r = INR xy:RH13:D y /\ SF ySF yun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RD run:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 lr:=MaxRlist l:RH11:exists n : nat, r = INR nx:natH12:r = INR xy:RH13:D y /\ SF yD yun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 lr:=MaxRlist l:RH11:exists n : nat, r = INR nx:natH12:r = INR xy:RH13:D y /\ SF ySF yun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RD run:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 lr:=MaxRlist l:RH11:exists n : nat, r = INR nx:natH12:r = INR xy:RH13:D y /\ SF ySF yun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RD run:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RD run:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> ind (subfamily f0 SF) rD run:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> D r /\ SF rIn r l -> D run:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> D r /\ SF rIn r lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> D r /\ SF rIn r lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> D r /\ SF r(exists z : R, intersection_domain (ind f0) SF z) -> exists y : R, In y lun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> D r /\ SF rexists z : R, intersection_domain (ind f0) SF zun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> D r /\ SF rexists z : R, intersection_domain (ind f0) SF zun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> D r /\ SF rH13:exists z : R, intersection_domain (ind f0) SF zexists z : R, intersection_domain (ind f0) SF zun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> D r /\ SF rH13:~ (exists z : R, intersection_domain (ind f0) SF z)exists z : R, intersection_domain (ind f0) SF zun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0H6:intersection_vide_in X f0H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)SF:R -> PropH8: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:RlistH9:forall x : R, ind (subfamily f0 SF) x <-> In x lr:=MaxRlist l:RH11:ind (subfamily f0 SF) r -> In r lH12:In r l -> D r /\ SF rH13:~ (exists z : R, intersection_domain (ind f0) SF z)exists z : R, intersection_domain (ind f0) SF zun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0intersection_vide_in X f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:Rind f0 x -> included (f0 x) Xun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:R~ (exists y : R, intersection_family f0 y)un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:RH6:D xincluded (adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x)) (adherence X)un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:RH6:D xincluded (adherence X) Xun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:R~ (exists y : R, intersection_family f0 y)un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:RH6:D xincluded (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x) Xun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:RH6:D xincluded (adherence X) Xun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:R~ (exists y : R, intersection_family f0 y)un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:RH6:D xincluded (adherence X) Xun:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:R~ (exists y : R, intersection_family f0 y)un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x0 : R => exists n : nat, x0 = INR n:R -> Propg:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> D x0f0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0x:R~ (exists y : R, intersection_family f0 y)un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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 f0unfold family_closed_set; unfold f0; simpl; unfold g; intro; apply adherence_P3. Qed. (********************************************************)un:nat -> RX:R -> PropH:compact XH0:forall n : nat, X (un n)H1:exists z : R, X zD:=fun x : R => exists n : nat, x = INR n:R -> Propg:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> PropH2:forall x : R, (exists y : R, g x y) -> D xf0:={| ind := D; f := g; cond_fam := H2 |}:familyH3: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
(********************************************************) 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 = yunfold is_lub; intros; elim H; elim H0; intros; apply Rle_antisym; [ apply (H4 _ H1) | apply (H2 _ H3) ]. Qed.forall (E : R -> Prop) (x y : R), is_lub E x -> is_lub E y -> x = yforall 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 -> PropH: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 -> PropH:~ (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 -> PropH:exists y : R, X yx:RH0:X xH1: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 -> PropH:exists y : R, X yx:RH0:X xH1:~ (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 -> PropH:~ (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 -> PropH:exists y : R, X yx:RH0:X xH1:exists y : R, X y /\ y <> xx0:RH2:X x0 /\ x0 <> xH3:X x0H4:x0 <> xX x /\ X x0 /\ x <> x0X:R -> PropH:exists y : R, X yx:RH0:X xH1:~ (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 -> PropH:~ (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 -> PropH:exists y : R, X yx:RH0:X xH1:~ (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 -> PropH:~ (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 -> PropH:exists y : R, X yx:RH0:X xH1:~ (exists y : R, X y /\ y <> x)X xX:R -> PropH:exists y : R, X yx:RH0:X xH1:~ (exists y : R, X y /\ y <> x)forall x0 : R, X x0 -> x0 = xX:R -> PropH:~ (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 -> PropH:exists y : R, X yx:RH0:X xH1:~ (exists y : R, X y /\ y <> x)forall x0 : R, X x0 -> x0 = xX:R -> PropH:~ (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 -> PropH:exists y : R, X yx:RH0:X xH1:~ (exists y : R, X y /\ y <> x)x0:RH2:X x0H3:x0 = xx0 = xX:R -> PropH:exists y : R, X yx:RH0:X xH1:~ (exists y : R, X y /\ y <> x)x0:RH2:X x0H3:x0 <> xx0 = xX:R -> PropH:~ (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 -> PropH:exists y : R, X yx:RH0:X xH1:~ (exists y : R, X y /\ y <> x)x0:RH2:X x0H3:x0 <> xx0 = xX:R -> PropH:~ (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.X:R -> PropH:~ (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)forall (f0 : R -> R) (X : R -> Prop), compact X -> (forall x : R, X x -> continuity_pt f0 x) -> uniform_continuity f0 Xforall (f0 : R -> R) (X : R -> Prop), compact X -> (forall x : R, X x -> continuity_pt f0 x) -> uniform_continuity f0 X(* X is empty *)f0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:~ (exists y : R, X y)uniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:(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 Xf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:(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 has only one element *)f0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists y : R, X y /\ (forall x : R, X x -> x = y)uniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yuniform_continuity f0 X(* X has at least two distinct elements *)f0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yuniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yexists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < Mf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yX_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < Muniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x1 : R, X x1 -> continuity_pt f0 x1Hyp:exists x1 y : R, X x1 /\ X y /\ x1 <> yH1:exists m M : R, forall x1 : R, X x1 -> m <= x1 <= Mx:RH2:exists M : R, forall x1 : R, X x1 -> x <= x1 <= Mx0:RH3:forall x1 : R, X x1 -> x <= x1 <= x0forall x1 : R, X x1 -> x <= x1 <= x0f0:R -> RX:R -> PropH0:compact XH:forall x1 : R, X x1 -> continuity_pt f0 x1Hyp:exists x1 y : R, X x1 /\ X y /\ x1 <> yH1:exists m M : R, forall x1 : R, X x1 -> m <= x1 <= Mx:RH2:exists M : R, forall x1 : R, X x1 -> x <= x1 <= Mx0:RH3:forall x1 : R, X x1 -> x <= x1 <= x0x < x0f0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yX_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < Muniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x1 : R, X x1 -> continuity_pt f0 x1Hyp:exists x1 y : R, X x1 /\ X y /\ x1 <> yH1:exists m M : R, forall x1 : R, X x1 -> m <= x1 <= Mx:RH2:exists M : R, forall x1 : R, X x1 -> x <= x1 <= Mx0:RH3:forall x1 : R, X x1 -> x <= x1 <= x0x < x0f0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yX_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < Muniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x3 : R, X x3 -> continuity_pt f0 x3Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> yH1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= Mx:RH2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= Mx0:RH3:forall x3 : R, X x3 -> x <= x3 <= x0x1:RH4:exists y : R, X x1 /\ X y /\ x1 <> yx2:RH5:X x1 /\ X x2 /\ x1 <> x2H6:X x1H8:X x2H9:x1 <> x2H10:x <= x1 <= x0H11:x <= x2 <= x0H7:x <= x1H12:x1 <= x0H13:x <= x2H14:x2 <= x0r:x < x0x < x0f0:R -> RX:R -> PropH0:compact XH:forall x3 : R, X x3 -> continuity_pt f0 x3Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> yH1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= Mx:RH2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= Mx0:RH3:forall x3 : R, X x3 -> x <= x3 <= x0x1:RH4:exists y : R, X x1 /\ X y /\ x1 <> yx2:RH5:X x1 /\ X x2 /\ x1 <> x2H6:X x1H8:X x2H9:x1 <> x2H10:x <= x1 <= x0H11:x <= x2 <= x0H7:x <= x1H12:x1 <= x0H13:x <= x2H14:x2 <= x0H15:x = x0x < x0f0:R -> RX:R -> PropH0:compact XH:forall x3 : R, X x3 -> continuity_pt f0 x3Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> yH1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= Mx:RH2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= Mx0:RH3:forall x3 : R, X x3 -> x <= x3 <= x0x1:RH4:exists y : R, X x1 /\ X y /\ x1 <> yx2:RH5:X x1 /\ X x2 /\ x1 <> x2H6:X x1H8:X x2H9:x1 <> x2H10:x <= x1 <= x0H11:x <= x2 <= x0H7:x <= x1H12:x1 <= x0H13:x <= x2H14:x2 <= x0H15:x > x0x < x0f0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yX_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < Muniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x3 : R, X x3 -> continuity_pt f0 x3Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> yH1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= Mx:RH2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= Mx0:RH3:forall x3 : R, X x3 -> x <= x3 <= x0x1:RH4:exists y : R, X x1 /\ X y /\ x1 <> yx2:RH5:X x1 /\ X x2 /\ x1 <> x2H6:X x1H8:X x2H9:x1 <> x2H10:x <= x1 <= x0H11:x <= x2 <= x0H7:x <= x1H12:x1 <= x0H13:x <= x2H14:x2 <= x0H15:x = x0x < x0f0:R -> RX:R -> PropH0:compact XH:forall x3 : R, X x3 -> continuity_pt f0 x3Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> yH1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= Mx:RH2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= Mx0:RH3:forall x3 : R, X x3 -> x <= x3 <= x0x1:RH4:exists y : R, X x1 /\ X y /\ x1 <> yx2:RH5:X x1 /\ X x2 /\ x1 <> x2H6:X x1H8:X x2H9:x1 <> x2H10:x <= x1 <= x0H11:x <= x2 <= x0H7:x <= x1H12:x1 <= x0H13:x <= x2H14:x2 <= x0H15:x > x0x < x0f0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yX_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < Muniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x3 : R, X x3 -> continuity_pt f0 x3Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> yH1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= Mx:RH2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= Mx0:RH3:forall x3 : R, X x3 -> x <= x3 <= x0x1:RH4:exists y : R, X x1 /\ X y /\ x1 <> yx2:RH5:X x1 /\ X x2 /\ x1 <> x2H6:X x1H8:X x2H9:x1 <> x2H10:x <= x1 <= x0H11:x <= x2 <= x0H7:x <= x1H12:x1 <= x0H13:x <= x2H14:x2 <= x0H15:x > x0x < x0f0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yX_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < Muniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xHyp:exists x y : R, X x /\ X y /\ x <> yX_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < Muniform_continuity f0 Xf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealforall t : posreal, 0 < t / 2f0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < epsf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < epsf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> Propexists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < epsf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> Propforall x : R, (exists y : R, g x y) -> X xf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xexists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < epsf0:R -> RX:R -> PropH0:compact XH:forall x : R, X x -> continuity_pt f0 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xexists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familycovering_open_set X f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familycovering X f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xX xf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xexists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xexists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> Propbound Ef0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound Eexists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound Eexists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound Eexists x1 : R, E x1f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 20 < Rmin x0 (M - m) <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 20 < Rmin x0 (M - m)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2Rmin x0 (M - m) <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2r:x0 <= M - m0 < x0f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2n:~ x0 <= M - m0 < M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2Rmin x0 (M - m) <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2n:~ x0 <= M - m0 < M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2Rmin x0 (M - m) <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2Rmin x0 (M - m) <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2z:RH8:Rabs (z - x) < Rmin x0 (M - m)H9:x = zRabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2z:RH8:Rabs (z - x) < Rmin x0 (M - m)H9:x <> zRabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2z:RH8:Rabs (z - x) < Rmin x0 (M - m)H9:x <> zRabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2z:RH8:Rabs (z - x) < Rmin x0 (M - m)H9:x <> zD_x no_cond x zf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2z:RH8:Rabs (z - x) < Rmin x0 (M - m)H9:x <> zRabs (z - x) < x0f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH6:bound EH5:x0 > 0H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2z:RH8:Rabs (z - x) < Rmin x0 (M - m)H9:x <> zRabs (z - x) < x0f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x1 : R, E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1exists 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 |} xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < 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 |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mforall z : R, Rabs (z - x) < {| pos := x1; cond_pos := H8 |} -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mz:RH10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}(exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mz:RH10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}exists alp : R, Rabs (z - x) < alp <= x1 /\ E alpf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x3 : R, E x3x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mz:RH10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}H11:exists alp : R, Rabs (z - x) < alp <= x1 /\ E alpx2:RH12:Rabs (z - x) < x2 <= x1H13:0 < x2 <= M - m /\ (forall z0 : R, Rabs (z0 - x) < x2 -> Rabs (f0 z0 - f0 x) < eps / 2)H14:0 < x2 <= M - mH15:forall z0 : R, Rabs (z0 - x) < x2 -> Rabs (f0 z0 - f0 x) < eps / 2Rabs (z - x) < x2f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mz:RH10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}exists alp : R, Rabs (z - x) < alp <= x1 /\ E alpf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mz:RH10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}exists alp : R, Rabs (z - x) < alp <= x1 /\ E alpf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mz:RH10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}H11:exists alp : R, Rabs (z - x) < alp <= x1 /\ E alpexists alp : R, Rabs (z - x) < alp <= x1 /\ E alpf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mz:RH10: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 alpf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mz:RH10: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 alpf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)H8:0 < x1H9:x1 <= M - mz:RH10: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) nH13:is_upper_bound E x1H14:forall b : R, is_upper_bound E b -> x1 <= bis_upper_bound E (Rabs (z - x)) -> exists alp : R, Rabs (z - x) < alp <= x1 /\ E alpf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)H8:0 < x1H9:x1 <= M - mz:RH10: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) nH13:is_upper_bound E x1H14:forall b : R, is_upper_bound E b -> x1 <= bis_upper_bound E (Rabs (z - x))f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)H8:0 < x1H9:x1 <= M - mz:RH10: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) nH13:is_upper_bound E x1H14:forall b : R, is_upper_bound E b -> x1 <= bis_upper_bound E (Rabs (z - x))f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x3 : R, E x3x1:Rp:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)H8:0 < x1H9:x1 <= M - mz:RH10: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) nH13:forall x3 : R, E x3 -> x3 <= x1H14:forall b : R, is_upper_bound E b -> x1 <= bx2:RH15:E x2H16:x2 <= x1r:x2 <= Rabs (z - x)x2 <= Rabs (z - x)f0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x3 : R, E x3x1:Rp:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)H8:0 < x1H9:x1 <= M - mz:RH10: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) n0H13:forall x3 : R, E x3 -> x3 <= x1H14:forall b : R, is_upper_bound E b -> x1 <= bx2:RH15:E x2H16:x2 <= x1n:~ x2 <= Rabs (z - x)x2 <= Rabs (z - x)f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x3 : R, E x3x1:Rp:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)H8:0 < x1H9:x1 <= M - mz:RH10: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) n0H13:forall x3 : R, E x3 -> x3 <= x1H14:forall b : R, is_upper_bound E b -> x1 <= bx2:RH15:E x2H16:x2 <= x1n:~ x2 <= Rabs (z - x)x2 <= Rabs (z - x)f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mis_lub E {| pos := x1; cond_pos := H8 |}f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mdisc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x1H8:0 < x1H9:x1 <= M - mdisc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} xf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x2 : R, E x2x1:Rp:is_lub E x10 < x1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x3 : R, E x3x1:Rp:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)x2:RH8:0 < x2 <= M - m /\ (forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2)H9:0 < x2 <= M - mH10:0 < x2H11:forall x3 : R, E x3 -> x3 <= x1H12:forall b : R, (forall x3 : R, E x3 -> x3 <= b) -> x1 <= b0 < x1f0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x3 : R, E x3x1:Rp:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)x2:RH8:0 < x2 <= M - m /\ (forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2)H9:0 < x2 <= M - mH10:0 < x2H11:forall x3 : R, E x3 -> x3 <= x1H12:forall b : R, (forall x3 : R, E x3 -> x3 <= b) -> x1 <= bx1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xH4: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:RH5: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 -> PropH6:bound EH7:exists x3 : R, E x3x1:Rp:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)x2:RH8:0 < x2 <= M - m /\ (forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2)H9:0 < x2 <= M - mH10:0 < x2H11:forall x3 : R, E x3 -> x3 <= x1H12:forall b : R, (forall x3 : R, E x3 -> x3 <= b) -> x1 <= bx1 <= M - mf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyfamily_open_set f'f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xopen_set (g x)f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0exists 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 -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0exists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0x2:RH7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2X xf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0x2:RH7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2exists 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 |} x2f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0exists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0x2:RH7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2exists 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 |} x2f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0exists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0x2:RH7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0x2:RH7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2is_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 |} x2f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0exists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0x2:RH7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2is_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 |} x2f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0exists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0x2:RH7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1f0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0x2:RH7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0exists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x = x0x2:RH7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0exists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0exists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):R0 < df0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dexists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dexists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:disc x0 {| pos := d; cond_pos := H7 |} x2X xf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:disc x0 {| pos := d; cond_pos := H7 |} x2exists 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 |} x2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:disc x0 {| pos := d; cond_pos := H7 |} x2exists 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 |} x2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:disc x0 {| pos := d; cond_pos := H7 |} x2forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:disc x0 {| pos := d; cond_pos := H7 |} x2is_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 |} x2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:disc x0 {| pos := d; cond_pos := H7 |} x2is_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 |} x2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:disc x0 {| pos := d; cond_pos := H7 |} x2H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1H10:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1f0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:disc x0 {| pos := d; cond_pos := H7 |} x2H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1H10:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:disc x0 {| pos := d; cond_pos := H7 |} x2H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1H10:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:Rabs (x2 - x0) < dH9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1H10:Rabs (x0 - x) < x1 / 2Rabs (x2 - x) <= Rabs (x2 - x0) + Rabs (x0 - x)f0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:Rabs (x2 - x0) < dH9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1H10:Rabs (x0 - x) < x1 / 2Rabs (x2 - x0) + Rabs (x0 - x) < x1 / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:Rabs (x2 - x0) < dH9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1H10:Rabs (x0 - x) < x1 / 2Rabs (x2 - x0) + Rabs (x0 - x) < x1 / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:X xx0:Rx1:posrealH4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2H5: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 |} x0H6:x <> x0d:=x1 / 2 - Rabs (x0 - x):RH7:0 < dx2:RH8:Rabs (x2 - x0) < dH9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1H10:Rabs (x0 - x) < x1 / 2Rabs (x2 - x0) + Rabs (x0 - x) < d + Rabs (x0 - x)f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (g x)f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xopen_set (fun _ : R => False)f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X x(fun _ : R => False) =_D g xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X x(fun _ : R => False) =_D g xf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xforall x0 : R, False -> g x x0f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xforall x0 : R, g x x0 -> Falsef0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyx:RH3:~ X xforall x0 : R, g x x0 -> Falsef0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3: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) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5: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) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2 -> exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiIn xi lf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lRabs (f0 x - f0 y) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lRabs (f0 x - f0 y) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 x - f0 y) <= Rabs (f0 x - f0 xi) + Rabs (f0 xi - f0 y)f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 x - f0 xi) + Rabs (f0 xi - f0 y) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 x - f0 xi) + Rabs (f0 xi - f0 y) < epsf0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 x - f0 xi) < eps / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 xi - f0 y) < eps / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2Rabs (x - xi) < pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2pos_Rl l' i / 2 < pos_Rl l' if0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 xi - f0 y) < eps / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2g xi xf0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2pos_Rl l' i / 2 < pos_Rl l' if0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 xi - f0 y) < eps / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2pos_Rl l' i / 2 < pos_Rl l' if0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 xi - f0 y) < eps / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 20 < 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 22 * (pos_Rl l' i * / 2) < 2 * pos_Rl l' if0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 xi - f0 y) < eps / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 22 * (pos_Rl l' i * / 2) < 2 * pos_Rl l' if0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 xi - f0 y) < eps / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2pos_Rl l' i * 1 < 2 * pos_Rl l' if0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 22 <> 0f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 xi - f0 y) < eps / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 22 <> 0f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 xi - f0 y) < eps / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iRabs (f0 xi - f0 y) < eps / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (y - xi) <= Rabs (y - x) + Rabs (x - xi)f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (y - x) + Rabs (x - xi) < pos_Rl l' if0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (y - x) + Rabs (x - xi) < pos_Rl l' if0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (y - x) < pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (x - xi) < pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (y - x) < D / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2D / 2 <= pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (x - xi) < pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2D / 2 <= pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (x - xi) < pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 20 <= / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2D <= pos_Rl l' if0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (x - xi) < pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2D <= pos_Rl l' if0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (x - xi) < pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0l:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lH6: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':RlistH7: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':RH9:0 < D / 2x, y:RH10:X xH11:X yH12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}xi:RH13:g xi x /\ DF xiH14:In xi lH15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0i:natH16:(i < Rlength l)%nat /\ xi = pos_Rl l iH17:(i < Rlength l)%natH18:xi = pos_Rl l iH19:0 < pos_Rl l' iH20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2H22:g xi xH23:DF xiH24:Rabs (x - xi) < pos_Rl l' i / 2Rabs (x - xi) < pos_Rl l' i / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lH6: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':RlistH7: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':R0 < D / 2f0:R -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 xm, M:RX_enc:forall x : R, X x -> m <= x <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x : R, (exists y : R, g x y) -> X xf':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x : R, X x -> exists y : R, g y x /\ DF yl:RlistH5:forall x : R, X x /\ DF x <-> In x lforall 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Propbound Ef0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound Eexists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound Eexists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound Eexists x0 : R, E x0f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x0 : R, E x0exists 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 -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH13: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:RH12:x0 > 0H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 20 < Rmin x0 (M - m) <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH13: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:RH12:x0 > 0H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x0 : R, E x0exists 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 -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH13: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:RH12:x0 > 0H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x0 : R, E x0exists 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 -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH13: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:RH12:x0 > 0H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2z:RH15:Rabs (z - x) < Rmin x0 (M - m)H16:x = zRabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH13: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:RH12:x0 > 0H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2z:RH15:Rabs (z - x) < Rmin x0 (M - m)H16:x <> zRabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x0 : R, E x0exists 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 -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH13: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:RH12:x0 > 0H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2z:RH15:Rabs (z - x) < Rmin x0 (M - m)H16:x <> zRabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x0 : R, E x0exists 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 -> RX:R -> PropH0: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 x0m, M:RX_enc:forall x0 : R, X x0 -> m <= x0 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x0 : R, (exists y : R, g x0 y) -> X x0f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF yl:RlistH5:forall x0 : R, X x0 /\ DF x0 <-> In x0 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x0 : R, E x0exists 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 -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0exists 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 -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < 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 -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - m0 < x0f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14: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 -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14: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 -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mforall z : R, Rabs (z - x) < x0 -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0(exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp) -> Rabs (f0 z - f0 x) < eps / 2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0exists alp : R, Rabs (z - x) < alp <= x0 /\ E alpf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0exists alp : R, Rabs (z - x) < alp <= x0 /\ E alpf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:exists alp : R, Rabs (z - x) < alp <= x0 /\ E alpexists alp : R, Rabs (z - x) < alp <= x0 /\ E alpf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)exists alp : R, Rabs (z - x) < alp <= x0 /\ E alpf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)exists alp : R, Rabs (z - x) < alp <= x0 /\ E alpf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)H17:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) nH18:is_upper_bound E x0H19:forall b : R, is_upper_bound E b -> x0 <= bis_upper_bound E (Rabs (z - x)) -> exists alp : R, Rabs (z - x) < alp <= x0 /\ E alpf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)H17:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) nH18:is_upper_bound E x0H19:forall b : R, is_upper_bound E b -> x0 <= bis_upper_bound E (Rabs (z - x))f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)H17:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) nH18:is_upper_bound E x0H19:forall b : R, is_upper_bound E b -> x0 <= bis_upper_bound E (Rabs (z - x))f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF yl:RlistH5:forall x2 : R, X x2 /\ DF x2 <-> In x2 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x2 : R, E x2x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)H17:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) nH18:forall x2 : R, E x2 -> x2 <= x0H19:forall b : R, is_upper_bound E b -> x0 <= bx1:RH20:E x1H21:x1 <= x0r:x1 <= Rabs (z - x)x1 <= Rabs (z - x)f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF yl:RlistH5:forall x2 : R, X x2 /\ DF x2 <-> In x2 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x2 : R, E x2x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0H18:forall x2 : R, E x2 -> x2 <= x0H19:forall b : R, is_upper_bound E b -> x0 <= bx1:RH20:E x1H21:x1 <= x0n:~ x1 <= Rabs (z - x)x1 <= Rabs (z - x)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF yl:RlistH5:forall x2 : R, X x2 /\ DF x2 <-> In x2 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x2 : R, E x2x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0H18:forall x2 : R, E x2 -> x2 <= x0H19:forall b : R, is_upper_bound E b -> x0 <= bx1:RH20:E x1H21:x1 <= x0n:~ x1 <= Rabs (z - x)x1 <= Rabs (z - x)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF yl:RlistH5:forall x2 : R, X x2 /\ DF x2 <-> In x2 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x2 : R, E x2x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0H18:forall x2 : R, E x2 -> x2 <= x0H19:forall b : R, is_upper_bound E b -> x0 <= bx1:RH20:E x1H21:x1 <= x0n:~ x1 <= Rabs (z - x)Rabs (z - x) < x1 <= x0f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF yl:RlistH5:forall x2 : R, X x2 /\ DF x2 <-> In x2 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x2 : R, E x2x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0H18:forall x2 : R, E x2 -> x2 <= x0H19:forall b : R, is_upper_bound E b -> x0 <= bx1:RH20:E x1H21:x1 <= x0n:~ x1 <= Rabs (z - x)E x1f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF yl:RlistH5:forall x2 : R, X x2 /\ DF x2 <-> In x2 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x2 : R, E x2x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)H13:0 < x0H14:x0 <= M - mz:RH15:Rabs (z - x) < x0H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0H18:forall x2 : R, E x2 -> x2 <= x0H19:forall b : R, is_upper_bound E b -> x0 <= bx1:RH20:E x1H21:x1 <= x0n:~ x1 <= Rabs (z - x)E x1f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mincluded (g x) (fun z : R => Rabs (z - x) < x0 / 2)f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF yl:RlistH5:forall x3 : R, X x3 /\ DF x3 <-> In x3 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x3 : R, E x3x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mx1:RH15: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 xH17: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 |} x1x2:posrealH18:(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 |} x1H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1x0 = x2 -> Rabs (x1 - x) < x0 / 2f0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF yl:RlistH5:forall x3 : R, X x3 /\ DF x3 <-> In x3 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x3 : R, E x3x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mx1:RH15: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 xH17: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 |} x1x2:posrealH18:(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 |} x1H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1x0 = x2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF yl:RlistH5:forall x3 : R, X x3 /\ DF x3 <-> In x3 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x3 : R, E x3x0:Rp:is_lub E x0H13:0 < x0H14:x0 <= M - mx1:RH15: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 xH17: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 |} x1x2:posrealH18:(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 |} x1H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1x0 = x2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF yl:RlistH5:forall x3 : R, X x3 /\ DF x3 <-> In x3 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x3 : R, E x3x0:Rp:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x0H13:0 < x0H14:x0 <= M - mx1:RH15: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 xH17: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 |} x1x2:posrealH18:(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 |} x1H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1is_lub ?E x0f0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF yl:RlistH5:forall x3 : R, X x3 /\ DF x3 <-> In x3 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x3 : R, E x3x0:Rp:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x0H13:0 < x0H14:x0 <= M - mx1:RH15: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 xH17: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 |} x1x2:posrealH18:(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 |} x1H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1is_lub ?E x2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x3m, M:RX_enc:forall x3 : R, X x3 -> m <= x3 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x3 : R, (exists y : R, g x3 y) -> X x3f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF yl:RlistH5:forall x3 : R, X x3 /\ DF x3 <-> In x3 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x3 : R, E x3x0:Rp:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x0H13:0 < x0H14:x0 <= M - mx1:RH15: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 xH17: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 |} x1x2:posrealH18:(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 |} x1H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2f0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x1m, M:RX_enc:forall x1 : R, X x1 -> m <= x1 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x1 : R, (exists y : R, g x1 y) -> X x1f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF yl:RlistH5:forall x1 : R, X x1 /\ DF x1 <-> In x1 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x1 : R, E x1x0:Rp:is_lub E x00 < x0 <= M - mf0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF yl:RlistH5:forall x2 : R, X x2 /\ DF x2 <-> In x2 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x2 : R, E x2x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)x1:RH13:0 < x1 <= M - m /\ (forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2)H14:0 < x1 <= M - mH15:0 < x1H16:forall x2 : R, E x2 -> x2 <= x0H17:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x0 <= b0 < x0f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF yl:RlistH5:forall x2 : R, X x2 /\ DF x2 <-> In x2 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x2 : R, E x2x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)x1:RH13:0 < x1 <= M - m /\ (forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2)H14:0 < x1 <= M - mH15:0 < x1H16:forall x2 : R, E x2 -> x2 <= x0H17:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x0 <= bx0 <= M - mapply H17; intros; unfold E in H18; elim H18; intros; elim H19; intros; assumption. Qed.f0:R -> RX:R -> PropH0: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 x2m, M:RX_enc:forall x2 : R, X x2 -> m <= x2 <= MHyp:m < Meps:posrealH1:forall t : posreal, 0 < t / 2g:=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 -> PropH2:forall x2 : R, (exists y : R, g x2 y) -> X x2f':={| ind := X; f := g; cond_fam := H2 |}:familyH3:covering_open_set X f'DF:R -> PropH4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF yl:RlistH5:forall x2 : R, X x2 /\ DF x2 <-> In x2 lx:RH6:In x lH7:X x /\ DF x -> In x lH8:In x l -> X x /\ DF xH9:X xH10:DF xE:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> PropH11:bound EH12:exists x2 : R, E x2x0:Rp:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)x1:RH13:0 < x1 <= M - m /\ (forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2)H14:0 < x1 <= M - mH15:0 < x1H16:forall x2 : R, E x2 -> x2 <= x0H17:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x0 <= bx0 <= M - m