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 SeqSeries.
Require Import Rtrigo1.
Require Import Ranalysis.
Local Open Scope R_scope.

(*******************************************)
(*            Newton's Integral            *)
(*******************************************)

Definition Newton_integrable (f:R -> R) (a b:R) : Type :=
  { g:R -> R | antiderivative f g a b \/ antiderivative f g b a }.

Definition NewtonInt (f:R -> R) (a b:R) (pr:Newton_integrable f a b) : R :=
  let (g,_) := pr in g b - g a.

(* If f is differentiable, then f' is Newton integrable (Tautology ?) *)

forall (f : Differential) (a b : R), Newton_integrable (fun x : R => derive_pt f x (cond_diff f x)) a b

forall (f : Differential) (a b : R), Newton_integrable (fun x : R => derive_pt f x (cond_diff f x)) a b
intros f a b; unfold Newton_integrable; exists (d1 f); unfold antiderivative; intros; case (Rle_dec a b); intro; [ left; split; [ intros; exists (cond_diff f x); reflexivity | assumption ] | right; split; [ intros; exists (cond_diff f x); reflexivity | auto with real ] ]. Defined. (* By definition, we have the Fondamental Theorem of Calculus *)

forall (f : Differential) (a b : R), NewtonInt (fun x : R => derive_pt f x (cond_diff f x)) a b (FTCN_step1 f a b) = f b - f a

forall (f : Differential) (a b : R), NewtonInt (fun x : R => derive_pt f x (cond_diff f x)) a b (FTCN_step1 f a b) = f b - f a
intros; unfold NewtonInt; reflexivity. Qed. (* $\int_a^a f$ exists forall a:R and f:R->R *)

forall (f : R -> R) (a : R), Newton_integrable f a a

forall (f : R -> R) (a : R), Newton_integrable f a a
f:R -> R
a:R

forall x : R, a <= x <= a -> exists pr : derivable_pt (fct_cte (f a) * id) x, f x = derive_pt (fct_cte (f a) * id) x pr
f:R -> R
a:R
a <= a
f:R -> R
a, x:R
H:a <= x <= a

derivable_pt (fct_cte (f a) * id) x
f:R -> R
a, x:R
H:a <= x <= a
H1:derivable_pt (fct_cte (f a) * id) x
exists pr : derivable_pt (fct_cte (f a) * id) x, f x = derive_pt (fct_cte (f a) * id) x pr
f:R -> R
a:R
a <= a
f:R -> R
a, x:R
H:a <= x <= a

derivable_pt (fct_cte (f a)) x
f:R -> R
a, x:R
H:a <= x <= a
derivable_pt id x
f:R -> R
a, x:R
H:a <= x <= a
H1:derivable_pt (fct_cte (f a) * id) x
exists pr : derivable_pt (fct_cte (f a) * id) x, f x = derive_pt (fct_cte (f a) * id) x pr
f:R -> R
a:R
a <= a
f:R -> R
a, x:R
H:a <= x <= a

derivable_pt id x
f:R -> R
a, x:R
H:a <= x <= a
H1:derivable_pt (fct_cte (f a) * id) x
exists pr : derivable_pt (fct_cte (f a) * id) x, f x = derive_pt (fct_cte (f a) * id) x pr
f:R -> R
a:R
a <= a
f:R -> R
a, x:R
H:a <= x <= a
H1:derivable_pt (fct_cte (f a) * id) x

exists pr : derivable_pt (fct_cte (f a) * id) x, f x = derive_pt (fct_cte (f a) * id) x pr
f:R -> R
a:R
a <= a
f:R -> R
a, x:R
H:a <= x <= a
H1:derivable_pt (fct_cte (f a) * id) x

x = a
f:R -> R
a, x:R
H:a <= x <= a
H1:derivable_pt (fct_cte (f a) * id) x
H2:x = a
f x = derive_pt (fct_cte (f a) * id) x H1
f:R -> R
a:R
a <= a
f:R -> R
a, x:R
H:a <= x <= a
H1:derivable_pt (fct_cte (f a) * id) x
H2:x = a

f x = derive_pt (fct_cte (f a) * id) x H1
f:R -> R
a:R
a <= a
f:R -> R
a:R

a <= a
right; reflexivity. Qed. (* $\int_a^a f = 0$ *)

forall (f : R -> R) (a : R), NewtonInt f a a (NewtonInt_P1 f a) = 0

forall (f : R -> R) (a : R), NewtonInt f a a (NewtonInt_P1 f a) = 0
f:R -> R
a:R

(let (g, _) := NewtonInt_P1 f a in g a - g a) = 0
f:R -> R
a:R
g:R -> R

g a - g a = 0
now apply Rminus_diag_eq. Qed. (* If $\int_a^b f$ exists, then $\int_b^a f$ exists too *)

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

forall (f : R -> R) (a b : R), Newton_integrable f a b -> Newton_integrable f b a
unfold Newton_integrable; intros; elim X; intros g H; exists g; tauto. Defined. (* $\int_a^b f = -\int_b^a f$ *)

forall (f : R -> R) (a b : R) (pr : Newton_integrable f a b), NewtonInt f a b pr = - NewtonInt f b a (NewtonInt_P3 f a b pr)

forall (f : R -> R) (a b : R) (pr : Newton_integrable f a b), NewtonInt f a b pr = - NewtonInt f b a (NewtonInt_P3 f a b pr)
f:R -> R
a, b:R
x:R -> R
H:antiderivative f x a b \/ antiderivative f x b a

NewtonInt f a b (exist (fun g : R -> R => antiderivative f g a b \/ antiderivative f g b a) x H) = - NewtonInt f b a (NewtonInt_P3 f a b (exist (fun g : R -> R => antiderivative f g a b \/ antiderivative f g b a) x H))
unfold NewtonInt, NewtonInt_P3; simpl; ring. Qed. (* The set of Newton integrable functions is a vectorial space *)

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

forall (f g : R -> R) (l a b : R), Newton_integrable f a b -> Newton_integrable g a b -> Newton_integrable (fun x : R => l * f x + g x) a b
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b
H0:antiderivative g x0 a b

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:a <= b
H0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr

(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr) /\ a <= b
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:a <= b
H0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr

forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:a <= b
H0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
a <= b
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 pr
H1:a <= b
H0:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 pr
x1:R
H2:a <= x1 <= b
x2:derivable_pt x0 x1
H3:g x1 = derive_pt x0 x1 x2
x3:derivable_pt x x1
H4:f x1 = derive_pt x x1 x3

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:a <= b
H0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
a <= b
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 pr
H1:a <= b
H0:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 pr
x1:R
H2:a <= x1 <= b
x2:derivable_pt x0 x1
H3:g x1 = derive_pt x0 x1 x2
x3:derivable_pt x x1
H4:f x1 = derive_pt x x1 x3

derivable_pt (fun y : R => l * x y + x0 y) x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 pr
H1:a <= b
H0:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 pr
x1:R
H2:a <= x1 <= b
x2:derivable_pt x0 x1
H3:g x1 = derive_pt x0 x1 x2
x3:derivable_pt x x1
H4:f x1 = derive_pt x x1 x3
H5:derivable_pt (fun y : R => l * x y + x0 y) x1
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:a <= b
H0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
a <= b
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 pr
H1:a <= b
H0:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 pr
x1:R
H2:a <= x1 <= b
x2:derivable_pt x0 x1
H3:g x1 = derive_pt x0 x1 x2
x3:derivable_pt x x1
H4:f x1 = derive_pt x x1 x3
H5:derivable_pt (fun y : R => l * x y + x0 y) x1

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:a <= b
H0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
a <= b
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:a <= b
H0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr

a <= b
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x a b
H0:antiderivative g x0 b a

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a < b

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b

forall x1 : R, a <= x1 <= a -> exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a

x1 = a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a

a <= x1 <= b
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b

b <= x1 <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a

derivable_pt x x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1

derivable_pt x0 x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1

derivable_pt (fun y : R => l * x y + x0 y) x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1

l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1

derive_pt x0 x1 H13 = g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1
l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1

l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1

derive_pt x x1 H12 = f x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1
H16:derive_pt x x1 H12 = f x1
l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= b
H0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= a
H1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:b <= a
H3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:a <= b
H5:a = b
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1
H16:derive_pt x x1 H12 = f x1

l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= b
H0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= a
H1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:b <= a
H3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:a <= b
H5:a = b

a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 a b

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b < a

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a

forall x1 : R, a <= x1 <= a -> exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a

x1 = a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a

a <= x1 <= b
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b

b <= x1 <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a

derivable_pt x x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1

derivable_pt x0 x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1

derivable_pt (fun y : R => l * x y + x0 y) x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1

l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1

derive_pt x0 x1 H13 = g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1
l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1

l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1

derive_pt x x1 H12 = f x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1
H16:derive_pt x x1 H12 = f x1
l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= a
H0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr
H2:a <= b
H3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr
H4:b <= a
H5:b = a
x1:R
H6:a <= x1 <= a
H7:a <= x1
H8:x1 <= a
H9:x1 = a
H10:a <= x1 <= b
H11:b <= x1 <= a
H12:derivable_pt x x1
H13:derivable_pt x0 x1
H14:derivable_pt (fun y : R => l * x y + x0 y) x1
H15:derive_pt x0 x1 H13 = g x1
H16:derive_pt x x1 H12 = f x1

l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a
a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= a
H0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= b
H1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
H2:a <= b
H3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H4:b <= a
H5:b = a

a <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a
antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:antiderivative f x b a
H0:antiderivative g x0 b a

antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) a b \/ antiderivative (fun x1 : R => l * f x1 + g x1) (fun y : R => l * x y + x0 y) b a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:b <= a
H0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr

forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:b <= a
H0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
b <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 pr
H1:b <= a
H0:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 pr
x1:R
H2:b <= x1 <= a
x2:derivable_pt x0 x1
H3:g x1 = derive_pt x0 x1 x2
x3:derivable_pt x x1
H4:f x1 = derive_pt x x1 x3

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:b <= a
H0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
b <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 pr
H1:b <= a
H0:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 pr
x1:R
H2:b <= x1 <= a
x2:derivable_pt x0 x1
H3:g x1 = derive_pt x0 x1 x2
x3:derivable_pt x x1
H4:f x1 = derive_pt x x1 x3

derivable_pt (fun y : R => l * x y + x0 y) x1
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 pr
H1:b <= a
H0:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 pr
x1:R
H2:b <= x1 <= a
x2:derivable_pt x0 x1
H3:g x1 = derive_pt x0 x1 x2
x3:derivable_pt x x1
H4:f x1 = derive_pt x x1 x3
H5:derivable_pt (fun y : R => l * x y + x0 y) x1
exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:b <= a
H0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
b <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 pr
H1:b <= a
H0:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 pr
x1:R
H2:b <= x1 <= a
x2:derivable_pt x0 x1
H3:g x1 = derive_pt x0 x1 x2
x3:derivable_pt x x1
H4:f x1 = derive_pt x x1 x3
H5:derivable_pt (fun y : R => l * x y + x0 y) x1

exists pr : derivable_pt (fun y : R => l * x y + x0 y) x1, l * f x1 + g x1 = derive_pt (fun y : R => l * x y + x0 y) x1 pr
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:b <= a
H0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr
b <= a
f, g:R -> R
l, a, b:R
X:{g0 : R -> R | antiderivative f g0 a b \/ antiderivative f g0 b a}
X0:{g0 : R -> R | antiderivative g g0 a b \/ antiderivative g g0 b a}
x:R -> R
p:antiderivative f x a b \/ antiderivative f x b a
x0:R -> R
p0:antiderivative g x0 a b \/ antiderivative g x0 b a
H:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr
H1:b <= a
H0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr

b <= a
assumption. Defined. (**********)

forall (f g F G : R -> R) (l a b : R), antiderivative f F a b -> antiderivative g G a b -> antiderivative (fun x : R => l * f x + g x) (fun x : R => l * F x + G x) a b

forall (f g F G : R -> R) (l a b : R), antiderivative f F a b -> antiderivative g G a b -> antiderivative (fun x : R => l * f x + g x) (fun x : R => l * F x + G x) a b
f, g, F, G:R -> R
l, a, b:R
H:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x pr
H0:a <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x pr
H2:a <= b

forall x : R, a <= x <= b -> exists pr : derivable_pt (fun x0 : R => l * F x0 + G x0) x, l * f x + g x = derive_pt (fun x0 : R => l * F x0 + G x0) x pr
f, g, F, G:R -> R
l, a, b:R
H:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x pr
H0:a <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x pr
H2:a <= b
a <= b
f, g, F, G:R -> R
l, a, b:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt G x2, g x2 = derive_pt G x2 pr
H0:a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F x2, f x2 = derive_pt F x2 pr
H2:a <= b
x:R
H3:a <= x <= b
x0:derivable_pt F x
H4:f x = derive_pt F x x0
x1:derivable_pt G x
H5:g x = derive_pt G x x1

exists pr : derivable_pt (fun x2 : R => l * F x2 + G x2) x, l * f x + g x = derive_pt (fun x2 : R => l * F x2 + G x2) x pr
f, g, F, G:R -> R
l, a, b:R
H:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x pr
H0:a <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x pr
H2:a <= b
a <= b
f, g, F, G:R -> R
l, a, b:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt G x2, g x2 = derive_pt G x2 pr
H0:a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F x2, f x2 = derive_pt F x2 pr
H2:a <= b
x:R
H3:a <= x <= b
x0:derivable_pt F x
H4:f x = derive_pt F x x0
x1:derivable_pt G x
H5:g x = derive_pt G x x1

derivable_pt (fun x2 : R => l * F x2 + G x2) x
f, g, F, G:R -> R
l, a, b:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt G x2, g x2 = derive_pt G x2 pr
H0:a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F x2, f x2 = derive_pt F x2 pr
H2:a <= b
x:R
H3:a <= x <= b
x0:derivable_pt F x
H4:f x = derive_pt F x x0
x1:derivable_pt G x
H5:g x = derive_pt G x x1
H6:derivable_pt (fun x2 : R => l * F x2 + G x2) x
exists pr : derivable_pt (fun x2 : R => l * F x2 + G x2) x, l * f x + g x = derive_pt (fun x2 : R => l * F x2 + G x2) x pr
f, g, F, G:R -> R
l, a, b:R
H:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x pr
H0:a <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x pr
H2:a <= b
a <= b
f, g, F, G:R -> R
l, a, b:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt G x2, g x2 = derive_pt G x2 pr
H0:a <= b
H1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F x2, f x2 = derive_pt F x2 pr
H2:a <= b
x:R
H3:a <= x <= b
x0:derivable_pt F x
H4:f x = derive_pt F x x0
x1:derivable_pt G x
H5:g x = derive_pt G x x1
H6:derivable_pt (fun x2 : R => l * F x2 + G x2) x

exists pr : derivable_pt (fun x2 : R => l * F x2 + G x2) x, l * f x + g x = derive_pt (fun x2 : R => l * F x2 + G x2) x pr
f, g, F, G:R -> R
l, a, b:R
H:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x pr
H0:a <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x pr
H2:a <= b
a <= b
f, g, F, G:R -> R
l, a, b:R
H:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x pr
H0:a <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x pr
H2:a <= b

a <= b
assumption. Qed. (* $\int_a^b \lambda f + g = \lambda \int_a^b f + \int_a^b f *)

forall (f g : R -> R) (l a b : R) (pr1 : Newton_integrable f a b) (pr2 : Newton_integrable g a b), NewtonInt (fun x : R => l * f x + g x) a b (NewtonInt_P5 f g l a b pr1 pr2) = l * NewtonInt f a b pr1 + NewtonInt g a b pr2

forall (f g : R -> R) (l a b : R) (pr1 : Newton_integrable f a b) (pr2 : Newton_integrable g a b), NewtonInt (fun x : R => l * f x + g x) a b (NewtonInt_P5 f g l a b pr1 pr2) = l * NewtonInt f a b pr1 + NewtonInt g a b pr2
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 a b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 a b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hlt:a < b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 a b
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a b
H3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2

a <= a <= b
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hlt:a < b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 a b
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a b
H3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2
H5:a <= a <= b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hlt:a < b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 a b
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a b
H3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2
H5:a <= a <= b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hlt:a < b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 a b
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a b
H3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2
H5:a <= a <= b

a <= b <= b
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hlt:a < b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 a b
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a b
H3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2
H5:a <= a <= b
H6:a <= b <= b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hlt:a < b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 a b
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a b
H3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2
H5:a <= a <= b
H6:a <= b <= b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 a b
H1:antiderivative g x1 b a

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b
H0:antiderivative f x0 b a

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hlt:a < b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Heq:a = b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x a b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
H0:antiderivative f x0 a b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
H0:antiderivative f x0 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
H0:antiderivative f x0 b a

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
H0:antiderivative f x0 b a
H1:antiderivative g x1 a b

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
H0:antiderivative f x0 b a
H1:antiderivative g x1 b a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b a
Hgt:a > b
H:antiderivative (fun x2 : R => l * f x2 + g x2) x b a
H0:antiderivative f x0 b a
H1:antiderivative g x1 b a

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hgt:a > b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x b a
H0:antiderivative f x0 b a
H1:antiderivative g x1 b a
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b a
H3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2

b <= a <= a
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hgt:a > b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x b a
H0:antiderivative f x0 b a
H1:antiderivative g x1 b a
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b a
H3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2
H5:b <= a <= a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hgt:a > b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x b a
H0:antiderivative f x0 b a
H1:antiderivative g x1 b a
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b a
H3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2
H5:b <= a <= a

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hgt:a > b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x b a
H0:antiderivative f x0 b a
H1:antiderivative g x1 b a
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b a
H3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2
H5:b <= a <= a

b <= b <= a
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hgt:a > b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x b a
H0:antiderivative f x0 b a
H1:antiderivative g x1 b a
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b a
H3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2
H5:b <= a <= a
H6:b <= b <= a
x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
f, g:R -> R
l, a, b:R
x0:R -> R
o0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
o1:antiderivative g x1 a b \/ antiderivative g x1 b a
x:R -> R
o:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b a
Hgt:a > b
H:antiderivative (fun x3 : R => l * f x3 + g x3) x b a
H0:antiderivative f x0 b a
H1:antiderivative g x1 b a
H2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b a
H3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + c
x2:R
H4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2
H5:b <= a <= a
H6:b <= b <= a

x b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)
assert (H7 := H4 _ H5); assert (H8 := H4 _ H6); rewrite H7; rewrite H8; ring. Qed.

forall (f F0 F1 : R -> R) (a b c : R), antiderivative f F0 a b -> antiderivative f F1 b c -> antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c

forall (f F0 F1 : R -> R) (a b c : R), antiderivative f F0 a b -> antiderivative f F1 b c -> antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H1:a <= b
H0:forall x : R, b <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H2:b <= c

forall x : R, a <= x <= c -> exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H1:a <= b
H0:forall x : R, b <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H2:b <= c
a <= c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H1:a <= b
H0:forall x : R, b <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H2:b <= c

forall x : R, a <= x <= c -> exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b

exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b

a <= x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b

exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0

exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0

derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0

forall eps : R, 0 < eps -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x1 (b - x):R

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x1 (b - x):R

0 < D
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x1 (b - x):R
r:x1 <= b - x

0 < x1
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x1 (b - x):R
n:~ x1 <= b - x
0 < b - x
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x1 (b - x):R
n:~ x1 <= b - x

0 < b - x
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}

Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b

Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - F0 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
r0:x + h <= b

Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
x + h <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
r0:x + h <= b

h <> 0
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
r0:x + h <= b
Rabs h < x1
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
x + h <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
r0:x + h <= b

Rabs h < x1
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
x + h <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b

x + h <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b

x + h < x + D
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
x + D <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b

h <= Rabs h
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
Rabs h < D
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
x + D <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b

Rabs h < D
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b
x + D <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
r:x <= b

x + D <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
eps:R
H9:0 < eps
H7:derive_pt F0 x x0 = f x
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x1 (b - x):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}

x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)

exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)

derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
H8:derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hlt:x < b
H5:a <= x <= b
x0:derivable_pt F0 x
H6:f x = derive_pt F0 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
H8:derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x

exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b

exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b

a <= x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b

exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b

b <= x <= c
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c

exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1

derive_pt F0 x x1 = f x
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x

exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x

derive_pt F1 x x0 = f x
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x

exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)

derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
x3:posreal
H15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x2 x3:R

0 < D
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
x3:posreal
H15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
x3:posreal
H15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x2 x3:R
r:x2 <= x3

0 < x2
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
x3:posreal
H15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x2 x3:R
n:~ x2 <= x3
0 < x3
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
x3:posreal
H15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
x3:posreal
H15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x2 x3:R
n:~ x2 <= x3

0 < x3
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
x3:posreal
H15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
x3:posreal
H15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b

Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - F0 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
r0:x + h <= b

Rabs ((F0 (x + h) - F0 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - F0 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
r0:x + h <= b

h <> 0
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
r0:x + h <= b
Rabs h < x3
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - F0 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
r0:x + h <= b

Rabs h < x3
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - F0 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b

Rabs ((F1 (x + h) + (F0 b - F1 b) - F0 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b

Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b
F1 (x + h) - F1 x = F1 (x + h) + (F0 b - F1 b) - F0 x
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b

h <> 0
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b
Rabs h < x2
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b
F1 (x + h) - F1 x = F1 (x + h) + (F0 b - F1 b) - F0 x
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b

Rabs h < x2
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b
F1 (x + h) - F1 x = F1 (x + h) + (F0 b - F1 b) - F0 x
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
r:x <= b
n:~ x + h <= b

F1 (x + h) - F1 x = F1 (x + h) + (F0 b - F1 b) - F0 x
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}
x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 pr
H1:a <= b
H0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps0
H12:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H13:0 < eps
x2:posreal
H14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
x3:posreal
H15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < eps
D:=Rmin x2 x3:R
H16:0 < D
h:R
H17:h <> 0
H18:Rabs h < {| pos := D; cond_pos := H16 |}

x <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)

exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)

derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
H14:derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x
exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Heq:x = b
H5:a <= x <= b
H6:b <= x <= c
x0:derivable_pt F1 x
H7:f x = derive_pt F1 x x0
x1:derivable_pt F0 x
H8:f x = derive_pt F0 x x1
H9:derive_pt F0 x x1 = f x
H10:derive_pt F1 x x0 = f x
H11:derivable_pt_lim F0 x (f x)
H12:derivable_pt_lim F1 x (f x)
H13:derivable_pt_lim (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x (f x)
H14:derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x

exists pr : derivable_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x, f x = derive_pt (fun x2 : R => if Rle_dec x2 b then F0 x2 else F1 x2 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b

exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b

b <= x <= c
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 pr
H1:a <= b
H0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c

exists pr : derivable_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x, f x = derive_pt (fun x0 : R => if Rle_dec x0 b then F0 x0 else F1 x0 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0

derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0

derive_pt F1 x x0 = f x
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
forall eps : R, 0 < eps -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x

forall eps : R, 0 < eps -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
D:=Rmin x1 (x - b):R

0 < D
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
D:=Rmin x1 (x - b):R
r:x1 <= x - b

0 < x1
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
D:=Rmin x1 (x - b):R
n:~ x1 <= x - b
0 < x - b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
D:=Rmin x1 (x - b):R
n:~ x1 <= x - b

0 < x - b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hle:x <= b

Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - F0 x) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b

Rabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b

Rabs ((F0 (x + h) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b

b < x + h -> Rabs ((F0 (x + h) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b
b < x + h
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b

b < x + h
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b

- h <= Rabs h
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b
Rabs h < x - b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b

Rabs h < x - b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b

Rabs h < D
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b
D <= x - b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hle':x + h <= b

D <= x - b
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b
Rabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b

Rabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < eps
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b

h <> 0
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b
Rabs h < x1
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b

Rabs h < x1
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b

Rabs h < D
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b
D <= x1
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 pr
H1:a <= b
H0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derive_pt F1 x x0 = f x
H8:forall eps0 : R, 0 < eps0 -> exists delta : posreal, forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps0
eps:R
H9:0 < eps
x1:posreal
H10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < eps
D:=Rmin x1 (x - b):R
H11:0 < D
h:R
H12:h <> 0
H13:Rabs h < {| pos := D; cond_pos := H11 |}
Hnle:~ x <= b
Hnle':~ x + h <= b

D <= x1
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)

exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)

derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
H8:derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x
exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 pr
H1:a <= b
H0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 pr
H2:b <= c
x:R
H3:a <= x
H4:x <= c
Hgt:x > b
H5:b <= x <= c
x0:derivable_pt F1 x
H6:f x = derive_pt F1 x x0
H7:derivable_pt_lim (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x (f x)
H8:derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x

exists pr : derivable_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x, f x = derive_pt (fun x1 : R => if Rle_dec x1 b then F0 x1 else F1 x1 + (F0 b - F1 b)) x pr
exists H8; symmetry ; apply derive_pt_eq_0; apply H7. Qed.

forall (f F0 F1 : R -> R) (a b c : R), antiderivative f F0 a b -> antiderivative f F1 c b -> antiderivative f F1 c a \/ antiderivative f F0 a c

forall (f F0 F1 : R -> R) (a b c : R), antiderivative f F0 a b -> antiderivative f F1 c b -> antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hle:a < c

antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:a = c
antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c
antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hle:a < c

forall x : R, a <= x <= c -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hle:a < c
a <= c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:a = c
antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c
antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hle:a < c

a <= c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:a = c
antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c
antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:a = c

antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c
antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:a = c

forall x : R, a <= x <= c -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:a = c
a <= c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c
antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:a = c

a <= c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c
antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c

antiderivative f F1 c a \/ antiderivative f F0 a c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c

forall x : R, c <= x <= a -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c
c <= a
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:c <= b
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:a > c

c <= a
left; assumption. Qed.

forall (f F0 F1 : R -> R) (a b c : R), antiderivative f F0 a b -> antiderivative f F1 a c -> antiderivative f F1 b c \/ antiderivative f F0 c b

forall (f F0 F1 : R -> R) (a b c : R), antiderivative f F0 a b -> antiderivative f F1 a c -> antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hlt:c < b

antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:c = b
antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b
antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hlt:c < b

forall x : R, c <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hlt:c < b
c <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:c = b
antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b
antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hlt:c < b

c <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:c = b
antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b
antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:c = b

antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b
antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:c = b

forall x : R, c <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:c = b
c <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b
antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Heq:c = b

c <= b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b
antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b

antiderivative f F1 b c \/ antiderivative f F0 c b
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b

forall x : R, b <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b
b <= c
f, F0, F1:R -> R
a, b, c:R
H:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x pr
H0:a <= c
H1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x pr
H2:a <= b
Hgt:c > b

b <= c
left; assumption. Qed.

forall (f : R -> R) (a b c : R), a < b -> b < c -> Newton_integrable f a b -> Newton_integrable f b c -> Newton_integrable f a c

forall (f : R -> R) (a b c : R), a < b -> b < c -> Newton_integrable f a b -> Newton_integrable f b c -> Newton_integrable f a c
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R

antiderivative f F0 a b
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R
antiderivative f F1 b c
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R
H:antiderivative f F0 a b

antiderivative f F0 a b
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R
H:antiderivative f F0 b a
antiderivative f F0 a b
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R
antiderivative f F1 b c
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R
H:antiderivative f F0 b a

antiderivative f F0 a b
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R
antiderivative f F1 b c
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R

antiderivative f F1 b c
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R
H:antiderivative f F1 b c

antiderivative f F1 b c
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R
H:antiderivative f F1 c b
antiderivative f F1 b c
f:R -> R
a, b, c:R
Hab:a < b
Hbc:b < c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
g:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> R
H:antiderivative f F1 c b

antiderivative f F1 b c
unfold antiderivative in H; elim H; clear H; intros; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 Hbc)). Qed.

forall (f : R -> R) (a b c : R), Newton_integrable f a b -> Newton_integrable f b c -> Newton_integrable f a c

forall (f : R -> R) (a b c : R), Newton_integrable f a b -> Newton_integrable f b c -> Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
(* a<b & b<c *)
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c

antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 a b

antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 b a
antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 a b
H2:antiderivative f F1 b c

antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 a b
H2:antiderivative f F1 c b
antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 b a
antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 a b
H2:antiderivative f F1 c b

antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 b a
antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 a b
H2:c <= b

antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 b a
antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:antiderivative f F0 b a

antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hlt':b < c
H:b <= a

antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) a c \/ antiderivative f (fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b)) c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Heq':b = c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
(* a<b & b=c *)
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
(* a<b & b>c *)
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c

antiderivative f F0 a c \/ antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 b c

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:b <= c

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c
H4:antiderivative f F1 c a

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c
H4:antiderivative f F0 a c
antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c
H4:c <= a

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c
H4:antiderivative f F0 a c
antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c
H4:antiderivative f F0 a c

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hlt'':a < c
H:antiderivative f F1 c b
H2:b <= a

antiderivative f F0 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Heq'':a = c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c

antiderivative f F1 a c \/ antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 b c

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:b <= c

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c
H4:antiderivative f F1 c a

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c
H4:antiderivative f F0 a c
antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c
H4:antiderivative f F0 a c

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 a b
H3:antiderivative f F1 c a \/ antiderivative f F0 a c
H4:a <= c

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a
antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:antiderivative f F0 b a

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hlt:a < b
Hgt':b > c
Hgt'':a > c
H:antiderivative f F1 c b
H2:b <= a

antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Heq:a = b

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Newton_integrable f a c
(* a=b *)
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
(* a>b & b<c *)
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c

antiderivative f F1 a c \/ antiderivative f F1 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
(*****************)
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:antiderivative f F0 a b

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:antiderivative f F0 b a
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:a <= b

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:antiderivative f F0 b a
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:antiderivative f F0 b a

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:antiderivative f F0 b a
H3:antiderivative f F1 a c \/ antiderivative f F0 c a

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:antiderivative f F0 b a
H3:antiderivative f F1 a c \/ antiderivative f F0 c a
H4:antiderivative f F1 a c

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:antiderivative f F0 b a
H3:antiderivative f F1 a c \/ antiderivative f F0 c a
H4:antiderivative f F0 c a
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:antiderivative f F0 b a
H3:antiderivative f F1 a c \/ antiderivative f F0 c a
H4:antiderivative f F0 c a

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 b c
H2:antiderivative f F0 b a
H3:antiderivative f F1 a c \/ antiderivative f F0 c a
H4:c <= a

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 c b
antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:antiderivative f F1 c b

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hlt'':a < c
H:c <= b

antiderivative f F1 a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Heq'':a = c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c

antiderivative f F0 a c \/ antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 a b

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:a <= b

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 b c

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 c b
antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 b c
H3:antiderivative f F1 a c \/ antiderivative f F0 c a

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 c b
antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 b c
H3:antiderivative f F1 a c \/ antiderivative f F0 c a
H4:antiderivative f F1 a c

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 b c
H3:antiderivative f F1 a c \/ antiderivative f F0 c a
H4:antiderivative f F0 c a
antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 c b
antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 b c
H3:antiderivative f F1 a c \/ antiderivative f F0 c a
H4:a <= c

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 b c
H3:antiderivative f F1 a c \/ antiderivative f F0 c a
H4:antiderivative f F0 c a
antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 c b
antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 b c
H3:antiderivative f F1 a c \/ antiderivative f F0 c a
H4:antiderivative f F0 c a

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 c b
antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:antiderivative f F1 c b

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hlt':b < c
Hgt'':a > c
H:antiderivative f F0 b a
H2:c <= b

antiderivative f F0 c a
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Heq':b = c

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
Newton_integrable f a c
(* a>b & b=c *)
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c

Newton_integrable f a c
(* a>b & b>c *)
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
X1:Newton_integrable f b a

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
X1:Newton_integrable f b a
X2:Newton_integrable f c b

Newton_integrable f a c
f:R -> R
a, b, c:R
X:Newton_integrable f a b
X0:Newton_integrable f b c
F0:R -> R
H0:antiderivative f F0 a b \/ antiderivative f F0 b a
F1:R -> R
H1:antiderivative f F1 b c \/ antiderivative f F1 c b
Hgt:a > b
Hgt':b > c
X1:Newton_integrable f b a
X2:Newton_integrable f c b

Newton_integrable f c a
apply NewtonInt_P7 with b; assumption. Qed. (* Chasles' relation *)

forall (f : R -> R) (a b c : R) (pr1 : Newton_integrable f a b) (pr2 : Newton_integrable f b c), NewtonInt f a c (NewtonInt_P8 f a b c pr1 pr2) = NewtonInt f a b pr1 + NewtonInt f b c pr2

forall (f : R -> R) (a b c : R) (pr1 : Newton_integrable f a b) (pr2 : Newton_integrable f b c), NewtonInt f a c (NewtonInt_P8 f a b c pr1 pr2) = NewtonInt f a b pr1 + NewtonInt f b c pr2
f:R -> R
a, b, c:R
pr1:Newton_integrable f a b
pr2:Newton_integrable f b c

(let (g, _) := NewtonInt_P8 f a b c pr1 pr2 in g c - g a) = (let (g, _) := pr1 in g b - g a) + (let (g, _) := pr2 in g c - g b)
f:R -> R
a, b, c:R
pr1:Newton_integrable f a b
pr2:Newton_integrable f b c
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a

x c - x a = (let (g, _) := pr1 in g b - g a) + (let (g, _) := pr2 in g c - g b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
pr2:Newton_integrable f b c
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a

x c - x a = x0 b - x0 a + (let (g, _) := pr2 in g c - g b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
(* a<b & b<c *)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x2 : R => if Rle_dec x2 b then x0 x2 else x1 x2 + (x0 b - x1 b)) a c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x2 : R => if Rle_dec x2 b then x0 x2 else x1 x2 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x2 : R, a <= x2 <= c -> x x2 = (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) x2 + c0

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2

a <= a <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c

a <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c

(if Rle_dec c b then x0 c else x1 c + (x0 b - x1 b)) + x2 - ((if Rle_dec a b then x0 a else x1 a + (x0 b - x1 b)) + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c
Hlea:a <= b

(if Rle_dec c b then x0 c else x1 c + (x0 b - x1 b)) + x2 - (x0 a + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c
Hnlea:~ a <= b
(if Rle_dec c b then x0 c else x1 c + (x0 b - x1 b)) + x2 - (x1 a + (x0 b - x1 b) + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c
Hlea:a <= b
Hlec:c <= b

x0 c + x2 - (x0 a + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c
Hlea:a <= b
Hnlec:~ c <= b
x1 c + (x0 b - x1 b) + x2 - (x0 a + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c
Hnlea:~ a <= b
(if Rle_dec c b then x0 c else x1 c + (x0 b - x1 b)) + x2 - (x1 a + (x0 b - x1 b) + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c
Hlea:a <= b
Hnlec:~ c <= b

x1 c + (x0 b - x1 b) + x2 - (x0 a + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c
Hnlea:~ a <= b
(if Rle_dec c b then x0 c else x1 c + (x0 b - x1 b)) + x2 - (x1 a + (x0 b - x1 b) + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a c
H3:exists c0 : R, forall x3 : R, a <= x3 <= c -> x x3 = (fun x4 : R => if Rle_dec x4 b then x0 x4 else x1 x4 + (x0 b - x1 b)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2
H5:a <= a <= c
H6:a <= c <= c
Hnlea:~ a <= b

(if Rle_dec c b then x0 c else x1 c + (x0 b - x1 b)) + x2 - (x1 a + (x0 b - x1 b) + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:antiderivative f x c a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 b c
H1:c <= a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:antiderivative f x1 c b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 a b
H0:c <= b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:antiderivative f x0 b a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hlt':b < c
H:b <= a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
(* a<b & b=c *)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c

x b - x a = x0 b - x0 a + (x1 b - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Heq':b = c

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x2 : R, a <= x2 <= b -> x x2 = x0 x2 + c0

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2

x0 b + x2 + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2
a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2

x0 b + x2 + - (x0 a + x2) = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2
a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2
a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2

a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2
a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x a b
H1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2

a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:antiderivative f x b a

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 a b
H0:b <= a

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:antiderivative f x0 b a

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hlt:a < b
Heq':b = c
H:b <= a

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
(* a<b & b>c *)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 b c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:b <= c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x2 : R => if Rle_dec x2 c then x x2 else x1 x2 + (x c - x1 c)) a b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x2 : R => if Rle_dec x2 c then x x2 else x1 x2 + (x c - x1 c)) a b
H3:exists c0 : R, forall x2 : R, a <= x2 <= b -> x0 x2 = (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) x2 + c0

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2

x c - x a = x0 b - ((if Rle_dec a c then x a else x1 a + (x c - x1 c)) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2

x c - x a = (if Rle_dec b c then x b else x1 b + (x c - x1 c)) + x2 - ((if Rle_dec a c then x a else x1 a + (x c - x1 c)) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
Hle:b <= c

x c - x a = x b + x2 - ((if Rle_dec a c then x a else x1 a + (x c - x1 c)) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
Hnle:~ b <= c
x c - x a = x1 b + (x c - x1 c) + x2 - ((if Rle_dec a c then x a else x1 a + (x c - x1 c)) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
Hnle:~ b <= c

x c - x a = x1 b + (x c - x1 c) + x2 - ((if Rle_dec a c then x a else x1 a + (x c - x1 c)) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
Hnle:~ b <= c
Hle':a <= c

x c - x a = x1 b + (x c - x1 c) + x2 - (x a + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
Hnle:~ b <= c
Hnle':~ a <= c
x c - x a = x1 b + (x c - x1 c) + x2 - (x1 a + (x c - x1 c) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
Hnle:~ b <= c
Hnle':~ a <= c

x c - x a = x1 b + (x c - x1 c) + x2 - (x1 a + (x c - x1 c) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2

a <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2
a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a b
H3:exists c0 : R, forall x3 : R, a <= x3 <= b -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x x4 else x1 x4 + (x c - x1 c)) x3 + c0
x2:R
H4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2

a <= a <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x2 : R => if Rle_dec x2 a then x x2 else x0 x2 + (x a - x0 a)) c b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x2 : R => if Rle_dec x2 a then x x2 else x0 x2 + (x a - x0 a)) c b
H3:exists c0 : R, forall x2 : R, c <= x2 <= b -> x1 x2 = (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) x2 + c0

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2

x c - x a = x0 b - x0 a + ((if Rle_dec c a then x c else x0 c + (x a - x0 a)) + x2 - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2

x c - x a = x0 b - x0 a + ((if Rle_dec c a then x c else x0 c + (x a - x0 a)) + x2 - ((if Rle_dec b a then x b else x0 b + (x a - x0 a)) + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
Hle:b <= a

x c - x a = x0 b - x0 a + ((if Rle_dec c a then x c else x0 c + (x a - x0 a)) + x2 - (x b + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
Hnle:~ b <= a
x c - x a = x0 b - x0 a + ((if Rle_dec c a then x c else x0 c + (x a - x0 a)) + x2 - (x0 b + (x a - x0 a) + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
Hnle:~ b <= a

x c - x a = x0 b - x0 a + ((if Rle_dec c a then x c else x0 c + (x a - x0 a)) + x2 - (x0 b + (x a - x0 a) + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
Hnle:~ b <= a
Hle':c <= a

x c - x a = x0 b - x0 a + (x c + x2 - (x0 b + (x a - x0 a) + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
Hnle:~ b <= a
c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
Hnle:~ b <= a

c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2

c <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 a b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c b
H3:exists c0 : R, forall x3 : R, c <= x3 <= b -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x x4 else x0 x4 + (x a - x0 a)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2

c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:antiderivative f x0 b a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hlt:a < b
Hgt':b > c
H:antiderivative f x1 c b
H0:b <= a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Heq:a = b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
(* a=b *)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x2 : R, b <= x2 <= c -> x x2 = x1 x2 + c0

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2

b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c

x c - (x1 b + x2) = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c

x1 c + x2 - (x1 b + x2) = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c
b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c

b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 b c
H1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2
H3:b <= c

b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b

b = c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
H1:b = c
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x b c
H0:antiderivative f x1 c b
H1:b = c

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 b c

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 b c

b = c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 b c
H1:b = c
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 b c
H1:b = c

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x2 : R, c <= x2 <= b -> x x2 = x1 x2 + c0

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2

c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b
x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b

x c - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b

x1 c + x2 - x b = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b

x1 c + x2 - (x1 b + x2) = x0 b - x0 b + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b
c <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b

c <= b <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b
c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x b c \/ antiderivative f x c b
Heq:a = b
H:antiderivative f x c b
H0:antiderivative f x1 c b
H1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0
x2:R
H2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2
H3:c <= b

c <= c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
(* a>b & b<c *)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 a b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:a <= b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x2 : R => if Rle_dec x2 a then x0 x2 else x x2 + (x0 a - x a)) b c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x2 : R => if Rle_dec x2 a then x0 x2 else x x2 + (x0 a - x a)) b c
H3:exists c0 : R, forall x2 : R, b <= x2 <= c -> x1 x2 = (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) x2 + c0

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2

x c - x a = x0 b - x0 a + (x1 c - ((if Rle_dec b a then x0 b else x b + (x0 a - x a)) + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2

x c - x a = x0 b - x0 a + ((if Rle_dec c a then x0 c else x c + (x0 a - x a)) + x2 - ((if Rle_dec b a then x0 b else x b + (x0 a - x a)) + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
r:b <= a

x c - x a = x0 b - x0 a + ((if Rle_dec c a then x0 c else x c + (x0 a - x a)) + x2 - (x0 b + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
r:b <= a
r0:c <= a

x c - x a = x0 b - x0 a + (x0 c + x2 - (x0 b + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
r:b <= a
n:~ c <= a
x c - x a = x0 b - x0 a + (x c + (x0 a - x a) + x2 - (x0 b + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
r:b <= a
r0:c <= a

a = c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
r:b <= a
r0:c <= a
H5:a = c
x c - x a = x0 b - x0 a + (x0 c + x2 - (x0 b + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
r:b <= a
n:~ c <= a
x c - x a = x0 b - x0 a + (x c + (x0 a - x a) + x2 - (x0 b + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
r:b <= a
r0:c <= a
H5:a = c

x c - x a = x0 b - x0 a + (x0 c + x2 - (x0 b + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
r:b <= a
n:~ c <= a
x c - x a = x0 b - x0 a + (x c + (x0 a - x a) + x2 - (x0 b + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
r:b <= a
n:~ c <= a

x c - x a = x0 b - x0 a + (x c + (x0 a - x a) + x2 - (x0 b + x2))
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2

b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2

b <= c <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2
b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x a c
H2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b c
H3:exists c0 : R, forall x3 : R, b <= x3 <= c -> x1 x3 = (fun x4 : R => if Rle_dec x4 a then x0 x4 else x x4 + (x0 a - x a)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2

b <= b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x2 : R => if Rle_dec x2 c then x1 x2 else x x2 + (x1 c - x c)) b a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x2 : R => if Rle_dec x2 c then x1 x2 else x x2 + (x1 c - x c)) b a
H3:exists c0 : R, forall x2 : R, b <= x2 <= a -> x0 x2 = (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) x2 + c0

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2

x c - x a = x0 b - ((if Rle_dec a c then x1 a else x a + (x1 c - x c)) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2

x c - x a = (if Rle_dec b c then x1 b else x b + (x1 c - x c)) + x2 - ((if Rle_dec a c then x1 a else x a + (x1 c - x c)) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
r:b <= c

x c - x a = x1 b + x2 - ((if Rle_dec a c then x1 a else x a + (x1 c - x c)) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
r:b <= c
r0:a <= c

x c - x a = x1 b + x2 - (x1 a + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
r:b <= c
n:~ a <= c
x c - x a = x1 b + x2 - (x a + (x1 c - x c) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
r:b <= c
r0:a <= c

a = c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
r:b <= c
r0:a <= c
H5:a = c
x c - x a = x1 b + x2 - (x1 a + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
r:b <= c
n:~ a <= c
x c - x a = x1 b + x2 - (x a + (x1 c - x c) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
r:b <= c
r0:a <= c
H5:a = c

x c - x a = x1 b + x2 - (x1 a + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
r:b <= c
n:~ a <= c
x c - x a = x1 b + x2 - (x a + (x1 c - x c) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
r:b <= c
n:~ a <= c

x c - x a = x1 b + x2 - (x a + (x1 c - x c) + x2) + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2

b <= c
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2

b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b a
H3:exists c0 : R, forall x3 : R, b <= x3 <= a -> x0 x3 = (fun x4 : R => if Rle_dec x4 c then x1 x4 else x x4 + (x1 c - x c)) x3 + c0
x2:R
H4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2

b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hlt':b < c
H:antiderivative f x0 b a
H0:c <= b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
(* a>b & b=c *)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c

x b - x a = x0 b - x0 a + (x1 b - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Heq':b = c

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 a b

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:a <= b

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x a b

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:a <= b

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x2 : R, b <= x2 <= a -> x x2 = x0 x2 + c0

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2

x b + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2

x0 b + x2 + - x a = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2

x0 b + x2 + - (x0 a + x2) = x0 b + - x0 a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2
b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2

b <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2
b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a b \/ antiderivative f x b a
Hgt:a > b
Heq':b = c
H:antiderivative f x0 b a
H0:antiderivative f x b a
H1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0
x2:R
H2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2

b <= b <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
(* a>b & b>c *)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 a b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:a <= b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 b c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:b <= c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x a c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:a <= c

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x2 : R => if Rle_dec x2 b then x1 x2 else x0 x2 + (x1 b - x0 b)) c a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x2 : R => if Rle_dec x2 b then x1 x2 else x0 x2 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x2 : R, c <= x2 <= a -> x x2 = (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) x2 + c0

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2

c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a

x c - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a

(if Rle_dec c b then x1 c else x0 c + (x1 b - x0 b)) + x2 - x a = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a

(if Rle_dec c b then x1 c else x0 c + (x1 b - x0 b)) + x2 - ((if Rle_dec a b then x1 a else x0 a + (x1 b - x0 b)) + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
Hle:a <= b

(if Rle_dec c b then x1 c else x0 c + (x1 b - x0 b)) + x2 - (x1 a + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
Hnle:~ a <= b
(if Rle_dec c b then x1 c else x0 c + (x1 b - x0 b)) + x2 - (x0 a + (x1 b - x0 b) + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
Hnle:~ a <= b

(if Rle_dec c b then x1 c else x0 c + (x1 b - x0 b)) + x2 - (x0 a + (x1 b - x0 b) + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
Hnle:~ a <= b
r:c <= b

x1 c + x2 - (x0 a + (x1 b - x0 b) + x2) = x0 b - x0 a + (x1 c - x1 b)
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
Hnle:~ a <= b
c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
Hnle:~ a <= b

c <= b
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a

c <= a <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a
c <= c <= a
f:R -> R
a, b, c:R
x0:R -> R
Hor0:antiderivative f x0 a b \/ antiderivative f x0 b a
x1:R -> R
Hor1:antiderivative f x1 b c \/ antiderivative f x1 c b
x:R -> R
Hor:antiderivative f x a c \/ antiderivative f x c a
Hgt:a > b
Hgt':b > c
H:antiderivative f x0 b a
H0:antiderivative f x1 c b
H1:antiderivative f x c a
H2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c a
H3:exists c0 : R, forall x3 : R, c <= x3 <= a -> x x3 = (fun x4 : R => if Rle_dec x4 b then x1 x4 else x0 x4 + (x1 b - x0 b)) x3 + c0
x2:R
H4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2
H5:c <= a

c <= c <= a
split; [ right; reflexivity | assumption ]. Qed.