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 bintros 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), Newton_integrable (fun x : R => derive_pt f x (cond_diff f x)) a bforall (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 aintros; unfold NewtonInt; reflexivity. Qed. (* $\int_a^a f$ exists forall a:R and f:R->R *)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 aforall (f : R -> R) (a : R), Newton_integrable f a aforall (f : R -> R) (a : R), Newton_integrable f a af:R -> Ra:Rforall 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 prf:R -> Ra:Ra <= af:R -> Ra, x:RH:a <= x <= aderivable_pt (fct_cte (f a) * id) xf:R -> Ra, x:RH:a <= x <= aH1:derivable_pt (fct_cte (f a) * id) xexists pr : derivable_pt (fct_cte (f a) * id) x, f x = derive_pt (fct_cte (f a) * id) x prf:R -> Ra:Ra <= af:R -> Ra, x:RH:a <= x <= aderivable_pt (fct_cte (f a)) xf:R -> Ra, x:RH:a <= x <= aderivable_pt id xf:R -> Ra, x:RH:a <= x <= aH1:derivable_pt (fct_cte (f a) * id) xexists pr : derivable_pt (fct_cte (f a) * id) x, f x = derive_pt (fct_cte (f a) * id) x prf:R -> Ra:Ra <= af:R -> Ra, x:RH:a <= x <= aderivable_pt id xf:R -> Ra, x:RH:a <= x <= aH1:derivable_pt (fct_cte (f a) * id) xexists pr : derivable_pt (fct_cte (f a) * id) x, f x = derive_pt (fct_cte (f a) * id) x prf:R -> Ra:Ra <= af:R -> Ra, x:RH:a <= x <= aH1:derivable_pt (fct_cte (f a) * id) xexists pr : derivable_pt (fct_cte (f a) * id) x, f x = derive_pt (fct_cte (f a) * id) x prf:R -> Ra:Ra <= af:R -> Ra, x:RH:a <= x <= aH1:derivable_pt (fct_cte (f a) * id) xx = af:R -> Ra, x:RH:a <= x <= aH1:derivable_pt (fct_cte (f a) * id) xH2:x = af x = derive_pt (fct_cte (f a) * id) x H1f:R -> Ra:Ra <= af:R -> Ra, x:RH:a <= x <= aH1:derivable_pt (fct_cte (f a) * id) xH2:x = af x = derive_pt (fct_cte (f a) * id) x H1f:R -> Ra:Ra <= aright; reflexivity. Qed. (* $\int_a^a f = 0$ *)f:R -> Ra:Ra <= aforall (f : R -> R) (a : R), NewtonInt f a a (NewtonInt_P1 f a) = 0forall (f : R -> R) (a : R), NewtonInt f a a (NewtonInt_P1 f a) = 0f:R -> Ra:R(let (g, _) := NewtonInt_P1 f a in g a - g a) = 0now apply Rminus_diag_eq. Qed. (* If $\int_a^b f$ exists, then $\int_b^a f$ exists too *)f:R -> Ra:Rg:R -> Rg a - g a = 0forall (f : R -> R) (a b : R), Newton_integrable f a b -> Newton_integrable f b aunfold 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), Newton_integrable f a b -> Newton_integrable f b aforall (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)unfold NewtonInt, NewtonInt_P3; simpl; ring. Qed. (* The set of Newton integrable functions is a vectorial space *)f:R -> Ra, b:Rx:R -> RH:antiderivative f x a b \/ antiderivative f x b aNewtonInt 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))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 bforall (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 bf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bH0:antiderivative g x0 a bantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:a <= bH0: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 <= bf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:a <= bH0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prforall 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:a <= bH0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pra <= bf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 prH1:a <= bH0:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 prx1:RH2:a <= x1 <= bx2:derivable_pt x0 x1H3:g x1 = derive_pt x0 x1 x2x3:derivable_pt x x1H4:f x1 = derive_pt x x1 x3exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:a <= bH0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pra <= bf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 prH1:a <= bH0:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 prx1:RH2:a <= x1 <= bx2:derivable_pt x0 x1H3:g x1 = derive_pt x0 x1 x2x3:derivable_pt x x1H4:f x1 = derive_pt x x1 x3derivable_pt (fun y : R => l * x y + x0 y) x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 prH1:a <= bH0:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 prx1:RH2:a <= x1 <= bx2:derivable_pt x0 x1H3:g x1 = derive_pt x0 x1 x2x3:derivable_pt x x1H4:f x1 = derive_pt x x1 x3H5:derivable_pt (fun y : R => l * x y + x0 y) x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:a <= bH0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pra <= bf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 prH1:a <= bH0:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 prx1:RH2:a <= x1 <= bx2:derivable_pt x0 x1H3:g x1 = derive_pt x0 x1 x2x3:derivable_pt x x1H4:f x1 = derive_pt x x1 x3H5:derivable_pt (fun y : R => l * x y + x0 y) x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:a <= bH0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pra <= bf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:a <= bH0:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pra <= bf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x a bH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a < bantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = bantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = bantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = bforall 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= ax1 = af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aa <= x1 <= bf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bb <= x1 <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aderivable_pt x x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1derivable_pt x0 x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1derivable_pt (fun y : R => l * x y + x0 y) x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1derive_pt x0 x1 H13 = g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1derive_pt x x1 H12 = f x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1H16:derive_pt x x1 H12 = f x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ a <= bH0:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ b <= aH1:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:b <= aH3:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:a <= bH5:a = bx1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1H16:derive_pt x x1 H12 = f x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ a <= bH0:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ b <= aH1:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:b <= aH3:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:a <= bH5:a = ba <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 a bantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b < aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aforall 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= ax1 = af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aa <= x1 <= bf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bb <= x1 <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aexists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aderivable_pt x x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1derivable_pt x0 x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1derivable_pt (fun y : R => l * x y + x0 y) x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1derive_pt x0 x1 H13 = g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1derive_pt x x1 H12 = f x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1H16:derive_pt x x1 H12 = f x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 pr) /\ b <= aH0:(forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 pr) /\ a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt x0 x2, g x2 = derive_pt x0 x2 prH2:a <= bH3:forall x2 : R, b <= x2 <= a -> exists pr : derivable_pt x x2, f x2 = derive_pt x x2 prH4:b <= aH5:b = ax1:RH6:a <= x1 <= aH7:a <= x1H8:x1 <= aH9:x1 = aH10:a <= x1 <= bH11:b <= x1 <= aH12:derivable_pt x x1H13:derivable_pt x0 x1H14:derivable_pt (fun y : R => l * x y + x0 y) x1H15:derive_pt x0 x1 H13 = g x1H16:derive_pt x x1 H12 = f x1l * derive_pt x x1 H12 + derive_pt x0 x1 H13 = l * f x1 + g x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:(forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 pr) /\ b <= aH0:(forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 pr) /\ a <= bH1:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prH2:a <= bH3:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH4:b <= aH5:b = aa <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:antiderivative f x b aH0:antiderivative g x0 b aantiderivative (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 af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:b <= aH0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prforall 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:b <= aH0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prb <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 prH1:b <= aH0:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 prx1:RH2:b <= x1 <= ax2:derivable_pt x0 x1H3:g x1 = derive_pt x0 x1 x2x3:derivable_pt x x1H4:f x1 = derive_pt x x1 x3exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:b <= aH0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prb <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 prH1:b <= aH0:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 prx1:RH2:b <= x1 <= ax2:derivable_pt x0 x1H3:g x1 = derive_pt x0 x1 x2x3:derivable_pt x x1H4:f x1 = derive_pt x x1 x3derivable_pt (fun y : R => l * x y + x0 y) x1f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 prH1:b <= aH0:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 prx1:RH2:b <= x1 <= ax2:derivable_pt x0 x1H3:g x1 = derive_pt x0 x1 x2x3:derivable_pt x x1H4:f x1 = derive_pt x x1 x3H5:derivable_pt (fun y : R => l * x y + x0 y) x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:b <= aH0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prb <= af, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x x4, f x4 = derive_pt x x4 prH1:b <= aH0:forall x4 : R, b <= x4 <= a -> exists pr : derivable_pt x0 x4, g x4 = derive_pt x0 x4 prx1:RH2:b <= x1 <= ax2:derivable_pt x0 x1H3:g x1 = derive_pt x0 x1 x2x3:derivable_pt x x1H4:f x1 = derive_pt x x1 x3H5:derivable_pt (fun y : R => l * x y + x0 y) x1exists 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 prf, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:b <= aH0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prb <= aassumption. Defined. (**********)f, g:R -> Rl, a, b:RX:{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 -> Rp:antiderivative f x a b \/ antiderivative f x b ax0:R -> Rp0:antiderivative g x0 a b \/ antiderivative g x0 b aH:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x x1, f x1 = derive_pt x x1 prH1:b <= aH0:forall x1 : R, b <= x1 <= a -> exists pr : derivable_pt x0 x1, g x1 = derive_pt x0 x1 prb <= aforall (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 bforall (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 bf, g, F, G:R -> Rl, a, b:RH:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x prH0:a <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x prH2:a <= bforall 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 prf, g, F, G:R -> Rl, a, b:RH:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x prH0:a <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x prH2:a <= ba <= bf, g, F, G:R -> Rl, a, b:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt G x2, g x2 = derive_pt G x2 prH0:a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F x2, f x2 = derive_pt F x2 prH2:a <= bx:RH3:a <= x <= bx0:derivable_pt F xH4:f x = derive_pt F x x0x1:derivable_pt G xH5:g x = derive_pt G x x1exists 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 prf, g, F, G:R -> Rl, a, b:RH:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x prH0:a <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x prH2:a <= ba <= bf, g, F, G:R -> Rl, a, b:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt G x2, g x2 = derive_pt G x2 prH0:a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F x2, f x2 = derive_pt F x2 prH2:a <= bx:RH3:a <= x <= bx0:derivable_pt F xH4:f x = derive_pt F x x0x1:derivable_pt G xH5:g x = derive_pt G x x1derivable_pt (fun x2 : R => l * F x2 + G x2) xf, g, F, G:R -> Rl, a, b:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt G x2, g x2 = derive_pt G x2 prH0:a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F x2, f x2 = derive_pt F x2 prH2:a <= bx:RH3:a <= x <= bx0:derivable_pt F xH4:f x = derive_pt F x x0x1:derivable_pt G xH5:g x = derive_pt G x x1H6:derivable_pt (fun x2 : R => l * F x2 + G x2) xexists 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 prf, g, F, G:R -> Rl, a, b:RH:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x prH0:a <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x prH2:a <= ba <= bf, g, F, G:R -> Rl, a, b:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt G x2, g x2 = derive_pt G x2 prH0:a <= bH1:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F x2, f x2 = derive_pt F x2 prH2:a <= bx:RH3:a <= x <= bx0:derivable_pt F xH4:f x = derive_pt F x x0x1:derivable_pt G xH5:g x = derive_pt G x x1H6:derivable_pt (fun x2 : R => l * F x2 + G x2) xexists 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 prf, g, F, G:R -> Rl, a, b:RH:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x prH0:a <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x prH2:a <= ba <= bassumption. Qed. (* $\int_a^b \lambda f + g = \lambda \int_a^b f + \int_a^b f *)f, g, F, G:R -> Rl, a, b:RH:forall x : R, a <= x <= b -> exists pr : derivable_pt G x, g x = derive_pt G x prH0:a <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F x, f x = derive_pt F x prH2:a <= ba <= bforall (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 pr2forall (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 pr2f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 a bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 a bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHlt:a < bH:antiderivative (fun x3 : R => l * f x3 + g x3) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 a bH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a bH3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2a <= a <= bf, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHlt:a < bH:antiderivative (fun x3 : R => l * f x3 + g x3) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 a bH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a bH3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2H5:a <= a <= bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHlt:a < bH:antiderivative (fun x3 : R => l * f x3 + g x3) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 a bH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a bH3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2H5:a <= a <= bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHlt:a < bH:antiderivative (fun x3 : R => l * f x3 + g x3) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 a bH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a bH3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2H5:a <= a <= ba <= b <= bf, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHlt:a < bH:antiderivative (fun x3 : R => l * f x3 + g x3) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 a bH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a bH3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2H5:a <= a <= bH6:a <= b <= bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHlt:a < bH:antiderivative (fun x3 : R => l * f x3 + g x3) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 a bH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) a bH3:exists c : R, forall x3 : R, a <= x3 <= b -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, a <= x3 <= b -> x x3 = l * x0 x3 + x1 x3 + x2H5:a <= a <= bH6:a <= b <= bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 a bH1:antiderivative g x1 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHlt:a < bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHeq:a = bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bH:antiderivative (fun x2 : R => l * f x2 + g x2) x a bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b aH0:antiderivative f x0 a bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b aH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b aH0:antiderivative f x0 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b aH0:antiderivative f x0 b aH1:antiderivative g x1 a bx b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b aH0:antiderivative f x0 b aH1:antiderivative g x1 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x2 : R => l * f x2 + g x2) x a b \/ antiderivative (fun x2 : R => l * f x2 + g x2) x b aHgt:a > bH:antiderivative (fun x2 : R => l * f x2 + g x2) x b aH0:antiderivative f x0 b aH1:antiderivative g x1 b ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHgt:a > bH:antiderivative (fun x3 : R => l * f x3 + g x3) x b aH0:antiderivative f x0 b aH1:antiderivative g x1 b aH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b aH3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2b <= a <= af, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHgt:a > bH:antiderivative (fun x3 : R => l * f x3 + g x3) x b aH0:antiderivative f x0 b aH1:antiderivative g x1 b aH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b aH3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2H5:b <= a <= ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHgt:a > bH:antiderivative (fun x3 : R => l * f x3 + g x3) x b aH0:antiderivative f x0 b aH1:antiderivative g x1 b aH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b aH3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2H5:b <= a <= ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHgt:a > bH:antiderivative (fun x3 : R => l * f x3 + g x3) x b aH0:antiderivative f x0 b aH1:antiderivative g x1 b aH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b aH3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2H5:b <= a <= ab <= b <= af, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHgt:a > bH:antiderivative (fun x3 : R => l * f x3 + g x3) x b aH0:antiderivative f x0 b aH1:antiderivative g x1 b aH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b aH3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2H5:b <= a <= aH6:b <= b <= ax 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.f, g:R -> Rl, a, b:Rx0:R -> Ro0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> Ro1:antiderivative g x1 a b \/ antiderivative g x1 b ax:R -> Ro:antiderivative (fun x3 : R => l * f x3 + g x3) x a b \/ antiderivative (fun x3 : R => l * f x3 + g x3) x b aHgt:a > bH:antiderivative (fun x3 : R => l * f x3 + g x3) x b aH0:antiderivative f x0 b aH1:antiderivative g x1 b aH2:antiderivative (fun x3 : R => l * f x3 + g x3) (fun x3 : R => l * x0 x3 + x1 x3) b aH3:exists c : R, forall x3 : R, b <= x3 <= a -> x x3 = (fun x4 : R => l * x0 x4 + x1 x4) x3 + cx2:RH4:forall x3 : R, b <= x3 <= a -> x x3 = l * x0 x3 + x1 x3 + x2H5:b <= a <= aH6:b <= b <= ax b - x a = l * (x0 b - x0 a) + (x1 b - x1 a)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 cforall (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 cf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH1:a <= bH0:forall x : R, b <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH2:b <= cforall 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 prf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH1:a <= bH0:forall x : R, b <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH2:b <= ca <= cf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH1:a <= bH0:forall x : R, b <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH2:b <= cforall 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < ba <= x <= bf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0exists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0derivable_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 -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0forall 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x1 (b - x):Rexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x1 (b - x):R0 < Df, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x1 (b - x):Rr:x1 <= b - x0 < x1f, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x1 (b - x):Rn:~ x1 <= b - x0 < b - xf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x1 (b - x):Rn:~ x1 <= b - x0 < b - xf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13: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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bRabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - F0 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= br0:x + h <= bRabs ((F0 (x + h) - F0 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bx + h <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= br0:x + h <= bh <> 0f, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= br0:x + h <= bRabs h < x1f, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bx + h <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= br0:x + h <= bRabs h < x1f, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bx + h <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bx + h <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bx + h < x + Df, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bx + D <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bh <= Rabs hf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bRabs h < Df, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bx + D <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bRabs h < Df, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bx + D <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}r:x <= bx + D <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0eps:RH9:0 < epsH7:derive_pt F0 x x0 = f xx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x1 (b - x):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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)) xf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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)) xexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHlt:x < bH5:a <= x <= bx0:derivable_pt F0 xH6:f x = derive_pt F0 x x0H7: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)) xexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = ba <= x <= bf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bb <= x <= cf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1derive_pt F0 x x1 = f xf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xderive_pt F1 x x0 = f xf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsx3:posrealH15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x2 x3:R0 < Df, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsx3:posrealH15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x2 x3:RH16:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsx3:posrealH15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x2 x3:Rr:x2 <= x30 < x2f, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsx3:posrealH15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x2 x3:Rn:~ x2 <= x30 < x3f, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsx3:posrealH15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x2 x3:RH16:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsx3:posrealH15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x2 x3:Rn:~ x2 <= x30 < x3f, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsx3:posrealH15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x2 x3:RH16:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h : R, h <> 0 -> Rabs h < x2 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsx3:posrealH15:forall h : R, h <> 0 -> Rabs h < x3 -> Rabs ((F0 (x + h) - F0 x) / h - f x) < epsD:=Rmin x2 x3:RH16:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bRabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - F0 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= br0:x + h <= bRabs ((F0 (x + h) - F0 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - F0 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= br0:x + h <= bh <> 0f, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= br0:x + h <= bRabs h < x3f, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - F0 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= br0:x + h <= bRabs h < x3f, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - F0 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - F0 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bRabs ((F1 (x + h) - F1 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bF1 (x + h) - F1 x = F1 (x + h) + (F0 b - F1 b) - F0 xf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bh <> 0f, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bRabs h < x2f, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bF1 (x + h) - F1 x = F1 (x + h) + (F0 b - F1 b) - F0 xf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bRabs h < x2f, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bF1 (x + h) - F1 x = F1 (x + h) + (F0 b - F1 b) - F0 xf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}r:x <= bn:~ x + h <= bF1 (x + h) - F1 x = F1 (x + h) + (F0 b - F1 b) - F0 xf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x4 : R, a <= x4 <= b -> exists pr : derivable_pt F0 x4, f x4 = derive_pt F0 x4 prH1:a <= bH0:forall x4 : R, b <= x4 <= c -> exists pr : derivable_pt F1 x4, f x4 = derive_pt F1 x4 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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) < eps0H12: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) < eps0eps:RH13:0 < epsx2:posrealH14:forall h0 : R, h0 <> 0 -> Rabs h0 < x2 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsx3:posrealH15:forall h0 : R, h0 <> 0 -> Rabs h0 < x3 -> Rabs ((F0 (x + h0) - F0 x) / h0 - f x) < epsD:=Rmin x2 x3:RH16:0 < Dh:RH17:h <> 0H18:Rabs h < {| pos := D; cond_pos := H16 |}x <= bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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)) xf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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)) xexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHeq:x = bH5:a <= x <= bH6:b <= x <= cx0:derivable_pt F1 xH7:f x = derive_pt F1 x x0x1:derivable_pt F0 xH8:f x = derive_pt F0 x x1H9:derive_pt F0 x x1 = f xH10:derive_pt F1 x x0 = f xH11: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)) xexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bb <= x <= cf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x0 : R, a <= x0 <= b -> exists pr : derivable_pt F0 x0, f x0 = derive_pt F0 x0 prH1:a <= bH0:forall x0 : R, b <= x0 <= c -> exists pr : derivable_pt F1 x0, f x0 = derive_pt F1 x0 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cexists 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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0derivable_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 -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0derive_pt F1 x x0 = f xf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xforall 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xforall 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsD:=Rmin x1 (x - b):R0 < Df, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsD:=Rmin x1 (x - b):Rr:x1 <= x - b0 < x1f, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsD:=Rmin x1 (x - b):Rn:~ x1 <= x - b0 < x - bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsD:=Rmin x1 (x - b):Rn:~ x1 <= x - b0 < x - bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h : R, h <> 0 -> Rabs h < x1 -> Rabs ((F1 (x + h) - F1 x) / h - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dexists 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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hle:x <= bRabs (((if Rle_dec (x + h) b then F0 (x + h) else F1 (x + h) + (F0 b - F1 b)) - F0 x) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bRabs (((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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bRabs (((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) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= bRabs ((F0 (x + h) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= bb < x + h -> Rabs ((F0 (x + h) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= bb < x + hf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= bb < x + hf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= b- h <= Rabs hf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= bRabs h < x - bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= bRabs h < x - bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= bRabs h < Df, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= bD <= x - bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHle':x + h <= bD <= x - bf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs ((F1 (x + h) + (F0 b - F1 b) - (F1 x + (F0 b - F1 b))) / h - f x) < epsf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bh <> 0f, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs h < x1f, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs h < x1f, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bRabs h < Df, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bD <= x1f, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x2 : R, a <= x2 <= b -> exists pr : derivable_pt F0 x2, f x2 = derive_pt F0 x2 prH1:a <= bH0:forall x2 : R, b <= x2 <= c -> exists pr : derivable_pt F1 x2, f x2 = derive_pt F1 x2 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7:derive_pt F1 x x0 = f xH8: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) < eps0eps:RH9:0 < epsx1:posrealH10:forall h0 : R, h0 <> 0 -> Rabs h0 < x1 -> Rabs ((F1 (x + h0) - F1 x) / h0 - f x) < epsD:=Rmin x1 (x - b):RH11:0 < Dh:RH12:h <> 0H13:Rabs h < {| pos := D; cond_pos := H11 |}Hnle:~ x <= bHnle':~ x + h <= bD <= x1f, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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 prf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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)) xf, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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)) xexists 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 prexists H8; symmetry ; apply derive_pt_eq_0; apply H7. Qed.f, F0, F1:R -> Ra, b, c:RH:forall x1 : R, a <= x1 <= b -> exists pr : derivable_pt F0 x1, f x1 = derive_pt F0 x1 prH1:a <= bH0:forall x1 : R, b <= x1 <= c -> exists pr : derivable_pt F1 x1, f x1 = derive_pt F1 x1 prH2:b <= cx:RH3:a <= xH4:x <= cHgt:x > bH5:b <= x <= cx0:derivable_pt F1 xH6:f x = derive_pt F1 x x0H7: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)) xexists 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 prforall (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 cforall (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 cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHle:a < cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:a = cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHle:a < cforall x : R, a <= x <= c -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHle:a < ca <= cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:a = cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHle:a < ca <= cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:a = cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:a = cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:a = cforall x : R, a <= x <= c -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:a = ca <= cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:a = ca <= cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cantiderivative f F1 c a \/ antiderivative f F0 a cf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cforall x : R, c <= x <= a -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prf, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cc <= aleft; assumption. Qed.f, F0, F1:R -> Ra, b, c:RH:forall x : R, c <= x <= b -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:c <= bH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:a > cc <= aforall (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 bforall (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 bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHlt:c < bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:c = bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHlt:c < bforall x : R, c <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHlt:c < bc <= bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:c = bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHlt:c < bc <= bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:c = bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:c = bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:c = bforall x : R, c <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:c = bc <= bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHeq:c = bc <= bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bantiderivative f F1 b c \/ antiderivative f F0 c bf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bforall x : R, b <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prf, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bb <= cleft; assumption. Qed.f, F0, F1:R -> Ra, b, c:RH:forall x : R, a <= x <= c -> exists pr : derivable_pt F1 x, f x = derive_pt F1 x prH0:a <= cH1:forall x : R, a <= x <= b -> exists pr : derivable_pt F0 x, f x = derive_pt F0 x prH2:a <= bHgt:c > bb <= cforall (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 cforall (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 cf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> Rantiderivative f F0 a bf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> Rantiderivative f F1 b cf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> RH:antiderivative f F0 a bantiderivative f F0 a bf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> RH:antiderivative f F0 b aantiderivative f F0 a bf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> Rantiderivative f F1 b cf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> RH:antiderivative f F0 b aantiderivative f F0 a bf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> Rantiderivative f F1 b cf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> Rantiderivative f F1 b cf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> RH:antiderivative f F1 b cantiderivative f F1 b cf:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> RH:antiderivative f F1 c bantiderivative f F1 b cunfold antiderivative in H; elim H; clear H; intros; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 Hbc)). Qed.f:R -> Ra, b, c:RHab:a < bHbc:b < cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bg:=fun x : R => if Rle_dec x b then F0 x else F1 x + (F0 b - F1 b):R -> RH:antiderivative f F1 c bantiderivative f F1 b cforall (f : R -> R) (a b c : R), Newton_integrable f a b -> Newton_integrable f b c -> Newton_integrable f a cforall (f : R -> R) (a b c : R), Newton_integrable f a b -> Newton_integrable f b c -> Newton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a c(* a<b & b<c *)f:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 a bantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 b aantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 a bH2:antiderivative f F1 b cantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 a bH2:antiderivative f F1 c bantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 b aantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 a bH2:antiderivative f F1 c bantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 b aantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 a bH2:c <= bantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 b aantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:antiderivative f F0 b aantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHlt':b < cH:b <= aantiderivative 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 af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a c(* a<b & b=c *)f:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a c(* a<b & b>c *)f:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cantiderivative f F0 a c \/ antiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 b cantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:b <= cantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 a bantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cH4:antiderivative f F1 c aantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cH4:antiderivative f F0 a cantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cH4:c <= aantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cH4:antiderivative f F0 a cantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cH4:antiderivative f F0 a cantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHlt'':a < cH:antiderivative f F1 c bH2:b <= aantiderivative f F0 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cantiderivative f F1 a c \/ antiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 b cantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:b <= cantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 a bantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cH4:antiderivative f F1 c aantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cH4:antiderivative f F0 a cantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cH4:antiderivative f F0 a cantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 a bH3:antiderivative f F1 c a \/ antiderivative f F0 a cH4:a <= cantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:antiderivative f F0 b aantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHlt:a < bHgt':b > cHgt'':a > cH:antiderivative f F1 c bH2:b <= aantiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a c(* a=b *)f:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHeq:a = bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bNewton_integrable f a c(* a>b & b<c *)f:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cantiderivative f F1 a c \/ antiderivative f F1 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a c(*****************)f:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 c bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:antiderivative f F0 a bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:antiderivative f F0 b aantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 c bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:a <= bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:antiderivative f F0 b aantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 c bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:antiderivative f F0 b aantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 c bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:antiderivative f F0 b aH3:antiderivative f F1 a c \/ antiderivative f F0 c aantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 c bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:antiderivative f F0 b aH3:antiderivative f F1 a c \/ antiderivative f F0 c aH4:antiderivative f F1 a cantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:antiderivative f F0 b aH3:antiderivative f F1 a c \/ antiderivative f F0 c aH4:antiderivative f F0 c aantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 c bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:antiderivative f F0 b aH3:antiderivative f F1 a c \/ antiderivative f F0 c aH4:antiderivative f F0 c aantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 c bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 b cH2:antiderivative f F0 b aH3:antiderivative f F1 a c \/ antiderivative f F0 c aH4:c <= aantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 c bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:antiderivative f F1 c bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHlt'':a < cH:c <= bantiderivative f F1 a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHeq'':a = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cantiderivative f F0 a c \/ antiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 a bantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:a <= bantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 b cantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 c bantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 b cH3:antiderivative f F1 a c \/ antiderivative f F0 c aantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 c bantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 b cH3:antiderivative f F1 a c \/ antiderivative f F0 c aH4:antiderivative f F1 a cantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 b cH3:antiderivative f F1 a c \/ antiderivative f F0 c aH4:antiderivative f F0 c aantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 c bantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 b cH3:antiderivative f F1 a c \/ antiderivative f F0 c aH4:a <= cantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 b cH3:antiderivative f F1 a c \/ antiderivative f F0 c aH4:antiderivative f F0 c aantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 c bantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 b cH3:antiderivative f F1 a c \/ antiderivative f F0 c aH4:antiderivative f F0 c aantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 c bantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:antiderivative f F1 c bantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHlt':b < cHgt'':a > cH:antiderivative f F0 b aH2:c <= bantiderivative f F0 c af:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a c(* a>b & b=c *)f:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHeq':b = cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a c(* a>b & b>c *)f:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cX1:Newton_integrable f b aNewton_integrable f a cf:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cX1:Newton_integrable f b aX2:Newton_integrable f c bNewton_integrable f a capply NewtonInt_P7 with b; assumption. Qed. (* Chasles' relation *)f:R -> Ra, b, c:RX:Newton_integrable f a bX0:Newton_integrable f b cF0:R -> RH0:antiderivative f F0 a b \/ antiderivative f F0 b aF1:R -> RH1:antiderivative f F1 b c \/ antiderivative f F1 c bHgt:a > bHgt':b > cX1:Newton_integrable f b aX2:Newton_integrable f c bNewton_integrable f c aforall (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 pr2forall (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 pr2f:R -> Ra, b, c:Rpr1:Newton_integrable f a bpr2: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 -> Ra, b, c:Rpr1:Newton_integrable f a bpr2:Newton_integrable f b cx:R -> RHor:antiderivative f x a c \/ antiderivative f x c ax c - x a = (let (g, _) := pr1 in g b - g a) + (let (g, _) := pr2 in g c - g b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b apr2:Newton_integrable f b cx:R -> RHor:antiderivative f x a c \/ antiderivative f x c ax c - x a = x0 b - x0 a + (let (g, _) := pr2 in g c - g b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)(* a<b & b<c *)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x2 : R => if Rle_dec x2 b then x0 x2 else x1 x2 + (x0 b - x1 b)) a cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x2 : R => if Rle_dec x2 b then x0 x2 else x1 x2 + (x0 b - x1 b)) a cH3: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 + c0x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2a <= a <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= ca <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6: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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cHlea: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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cHnlea:~ 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cHlea:a <= bHlec:c <= bx0 c + x2 - (x0 a + x2) = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cHlea:a <= bHnlec:~ c <= bx1 c + (x0 b - x1 b) + x2 - (x0 a + x2) = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cHnlea:~ 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cHlea:a <= bHnlec:~ c <= bx1 c + (x0 b - x1 b) + x2 - (x0 a + x2) = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cHnlea:~ 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) a cH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= c -> x x3 = (if Rle_dec x3 b then x0 x3 else x1 x3 + (x0 b - x1 b)) + x2H5:a <= a <= cH6:a <= c <= cHnlea:~ 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 b cH1:c <= ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 a bH0:c <= bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHlt':b < cH:b <= ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)(* a<b & b=c *)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx b - x a = x0 b - x0 a + (x1 b - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHeq':b = cx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x2 : R, a <= x2 <= b -> x x2 = x0 x2 + c0x b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2x b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2x0 b + x2 + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2x0 b + x2 + - (x0 a + x2) = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x a bH1:exists c0 : R, forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, a <= x3 <= b -> x x3 = x0 x3 + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 a bH0:b <= ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHlt:a < bHeq':b = cH:b <= ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)(* a<b & b>c *)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 b cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:b <= cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x2 : R => if Rle_dec x2 c then x x2 else x1 x2 + (x c - x1 c)) a bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x2 : R => if Rle_dec x2 c then x x2 else x1 x2 + (x c - x1 c)) a bH3: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 + c0x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2x 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2x 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2Hle:b <= cx 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2Hnle:~ b <= cx 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2Hnle:~ b <= cx 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2Hnle:~ b <= cHle':a <= cx c - x a = x1 b + (x c - x1 c) + x2 - (x a + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2Hnle:~ b <= cHnle':~ a <= cx c - x a = x1 b + (x c - x1 c) + x2 - (x1 a + (x c - x1 c) + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2Hnle:~ b <= cHnle':~ a <= cx c - x a = x1 b + (x c - x1 c) + x2 - (x1 a + (x c - x1 c) + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) a bH3: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 + c0x2:RH4:forall x3 : R, a <= x3 <= b -> x0 x3 = (if Rle_dec x3 c then x x3 else x1 x3 + (x c - x1 c)) + x2a <= a <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x2 : R => if Rle_dec x2 a then x x2 else x0 x2 + (x a - x0 a)) c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x2 : R => if Rle_dec x2 a then x x2 else x0 x2 + (x a - x0 a)) c bH3: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 + c0x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2x 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2x 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2Hle:b <= ax 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2Hnle:~ b <= ax 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2Hnle:~ b <= ax 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2Hnle:~ b <= aHle':c <= ax c - x a = x0 b - x0 a + (x c + x2 - (x0 b + (x a - x0 a) + x2))f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2Hnle:~ b <= ac <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2Hnle:~ b <= ac <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 a bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) c bH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= b -> x1 x3 = (if Rle_dec x3 a then x x3 else x0 x3 + (x a - x0 a)) + x2c <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHlt:a < bHgt':b > cH:antiderivative f x1 c bH0:b <= ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)(* a=b *)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHeq:a = bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x2 : R, b <= x2 <= c -> x x2 = x1 x2 + c0x c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2x c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cx c - (x1 b + x2) = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cb <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cx1 c + x2 - (x1 b + x2) = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cb <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cb <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cb <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cb <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 b cH1:exists c0 : R, forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= c -> x x3 = x1 x3 + x2H3:b <= cb <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bb = cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bH1:b = cx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x b cH0:antiderivative f x1 c bH1:b = cx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 b cx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 b cb = cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 b cH1:b = cx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 b cH1:b = cx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x2 : R, c <= x2 <= b -> x x2 = x1 x2 + c0x c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2x c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bx c - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bx1 c + x2 - x b = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bc <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bx1 c + x2 - (x1 b + x2) = x0 b - x0 b + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bc <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bc <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bc <= b <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bc <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x b c \/ antiderivative f x c bHeq:a = bH:antiderivative f x c bH0:antiderivative f x1 c bH1:exists c0 : R, forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + c0x2:RH2:forall x3 : R, c <= x3 <= b -> x x3 = x1 x3 + x2H3:c <= bc <= c <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)(* a>b & b<c *)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 a bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:a <= bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x2 : R => if Rle_dec x2 a then x0 x2 else x x2 + (x0 a - x a)) b cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x2 : R => if Rle_dec x2 a then x0 x2 else x x2 + (x0 a - x a)) b cH3: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 + c0x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2x 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2x 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2r:b <= ax 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2r:b <= ar0:c <= ax c - x a = x0 b - x0 a + (x0 c + x2 - (x0 b + x2))f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2r:b <= an:~ c <= ax c - x a = x0 b - x0 a + (x c + (x0 a - x a) + x2 - (x0 b + x2))f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2r:b <= ar0:c <= aa = cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2r:b <= ar0:c <= aH5:a = cx c - x a = x0 b - x0 a + (x0 c + x2 - (x0 b + x2))f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2r:b <= an:~ c <= ax c - x a = x0 b - x0 a + (x c + (x0 a - x a) + x2 - (x0 b + x2))f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2r:b <= ar0:c <= aH5:a = cx c - x a = x0 b - x0 a + (x0 c + x2 - (x0 b + x2))f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2r:b <= an:~ c <= ax c - x a = x0 b - x0 a + (x c + (x0 a - x a) + x2 - (x0 b + x2))f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2r:b <= an:~ c <= ax c - x a = x0 b - x0 a + (x c + (x0 a - x a) + x2 - (x0 b + x2))f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= c <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x a cH2:antiderivative f (fun x3 : R => if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) b cH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= c -> x1 x3 = (if Rle_dec x3 a then x0 x3 else x x3 + (x0 a - x a)) + x2b <= b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x2 : R => if Rle_dec x2 c then x1 x2 else x x2 + (x1 c - x c)) b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x2 : R => if Rle_dec x2 c then x1 x2 else x x2 + (x1 c - x c)) b aH3: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 + c0x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2x 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2x 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2r:b <= cx 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2r:b <= cr0:a <= cx c - x a = x1 b + x2 - (x1 a + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2r:b <= cn:~ a <= cx c - x a = x1 b + x2 - (x a + (x1 c - x c) + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2r:b <= cr0:a <= ca = cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2r:b <= cr0:a <= cH5:a = cx c - x a = x1 b + x2 - (x1 a + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2r:b <= cn:~ a <= cx c - x a = x1 b + x2 - (x a + (x1 c - x c) + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2r:b <= cr0:a <= cH5:a = cx c - x a = x1 b + x2 - (x1 a + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2r:b <= cn:~ a <= cx c - x a = x1 b + x2 - (x a + (x1 c - x c) + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2r:b <= cn:~ a <= cx c - x a = x1 b + x2 - (x a + (x1 c - x c) + x2) + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= cf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 b cH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) b aH3: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 + c0x2:RH4:forall x3 : R, b <= x3 <= a -> x0 x3 = (if Rle_dec x3 c then x1 x3 else x x3 + (x1 c - x c)) + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHlt':b < cH:antiderivative f x0 b aH0:c <= bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)(* a>b & b=c *)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx b - x a = x0 b - x0 a + (x1 b - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHeq':b = cx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 a bx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:a <= bx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x a bx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:a <= bx b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b ax b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x2 : R, b <= x2 <= a -> x x2 = x0 x2 + c0x b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2x b + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2x0 b + x2 + - x a = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2x0 b + x2 + - (x0 a + x2) = x0 b + - x0 af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2b <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a b \/ antiderivative f x b aHgt:a > bHeq':b = cH:antiderivative f x0 b aH0:antiderivative f x b aH1:exists c0 : R, forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + c0x2:RH2:forall x3 : R, b <= x3 <= a -> x x3 = x0 x3 + x2b <= b <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)(* a>b & b>c *)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 a bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:a <= bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 b cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:b <= cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x a cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:a <= cx c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x2 : R => if Rle_dec x2 b then x1 x2 else x0 x2 + (x1 b - x0 b)) c ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x2 : R => if Rle_dec x2 b then x1 x2 else x0 x2 + (x1 b - x0 b)) c aH3: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 + c0x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2x c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2c <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ax c - x a = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5: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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= c <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5: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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= c <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= aHle: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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= aHnle:~ 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= c <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= aHnle:~ 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 -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= c <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= aHnle:~ a <= br:c <= bx1 c + x2 - (x0 a + (x1 b - x0 b) + x2) = x0 b - x0 a + (x1 c - x1 b)f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= aHnle:~ a <= bc <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= c <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= aHnle:~ a <= bc <= bf:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= c <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= a <= af:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= c <= asplit; [ right; reflexivity | assumption ]. Qed.f:R -> Ra, b, c:Rx0:R -> RHor0:antiderivative f x0 a b \/ antiderivative f x0 b ax1:R -> RHor1:antiderivative f x1 b c \/ antiderivative f x1 c bx:R -> RHor:antiderivative f x a c \/ antiderivative f x c aHgt:a > bHgt':b > cH:antiderivative f x0 b aH0:antiderivative f x1 c bH1:antiderivative f x c aH2:antiderivative f (fun x3 : R => if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) c aH3: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 + c0x2:RH4:forall x3 : R, c <= x3 <= a -> x x3 = (if Rle_dec x3 b then x1 x3 else x0 x3 + (x1 b - x0 b)) + x2H5:c <= ac <= c <= a