Up to index of Isabelle/HOL/HOL-Complex
theory Integration(* ID : $Id: Integration.thy,v 1.34 2007/10/21 12:53:45 nipkow Exp $ Author : Jacques D. Fleuriot Copyright : 2000 University of Edinburgh Conversion to Isar and new proofs by Lawrence C Paulson, 2004 *) header{*Theory of Integration*} theory Integration imports MacLaurin begin text{*We follow John Harrison in formalizing the Gauge integral.*} definition --{*Partitions and tagged partitions etc.*} partition :: "[(real*real),nat => real] => bool" where "partition = (%(a,b) D. D 0 = a & (∃N. (∀n < N. D(n) < D(Suc n)) & (∀n ≥ N. D(n) = b)))" definition psize :: "(nat => real) => nat" where "psize D = (SOME N. (∀n < N. D(n) < D(Suc n)) & (∀n ≥ N. D(n) = D(N)))" definition tpart :: "[(real*real),((nat => real)*(nat =>real))] => bool" where "tpart = (%(a,b) (D,p). partition(a,b) D & (∀n. D(n) ≤ p(n) & p(n) ≤ D(Suc n)))" --{*Gauges and gauge-fine divisions*} definition gauge :: "[real => bool, real => real] => bool" where "gauge E g = (∀x. E x --> 0 < g(x))" definition fine :: "[real => real, ((nat => real)*(nat => real))] => bool" where "fine = (%g (D,p). ∀n. n < (psize D) --> D(Suc n) - D(n) < g(p n))" --{*Riemann sum*} definition rsum :: "[((nat=>real)*(nat=>real)),real=>real] => real" where "rsum = (%(D,p) f. ∑n=0..<psize(D). f(p n) * (D(Suc n) - D(n)))" --{*Gauge integrability (definite)*} definition Integral :: "[(real*real),real=>real,real] => bool" where "Integral = (%(a,b) f k. ∀e > 0. (∃g. gauge(%x. a ≤ x & x ≤ b) g & (∀D p. tpart(a,b) (D,p) & fine(g)(D,p) --> ¦rsum(D,p) f - k¦ < e)))" lemma partition_zero [simp]: "a = b ==> psize (%n. if n = 0 then a else b) = 0" by (auto simp add: psize_def) lemma partition_one [simp]: "a < b ==> psize (%n. if n = 0 then a else b) = 1" apply (simp add: psize_def) apply (rule some_equality, auto) apply (drule_tac x = 1 in spec, auto) done lemma partition_single [simp]: "a ≤ b ==> partition(a,b)(%n. if n = 0 then a else b)" by (auto simp add: partition_def order_le_less) lemma partition_lhs: "partition(a,b) D ==> (D(0) = a)" by (simp add: partition_def) lemma partition: "(partition(a,b) D) = ((D 0 = a) & (∀n < psize D. D n < D(Suc n)) & (∀n ≥ psize D. D n = b))" apply (simp add: partition_def, auto) apply (subgoal_tac [!] "psize D = N", auto) apply (simp_all (no_asm) add: psize_def) apply (rule_tac [!] some_equality, blast) prefer 2 apply blast apply (rule_tac [!] ccontr) apply (simp_all add: linorder_neq_iff, safe) apply (drule_tac x = Na in spec) apply (rotate_tac 3) apply (drule_tac x = "Suc Na" in spec, simp) apply (rotate_tac 2) apply (drule_tac x = N in spec, simp) apply (drule_tac x = Na in spec) apply (drule_tac x = "Suc Na" and P = "%n. Na≤n --> D n = D Na" in spec, auto) done lemma partition_rhs: "partition(a,b) D ==> (D(psize D) = b)" by (simp add: partition) lemma partition_rhs2: "[|partition(a,b) D; psize D ≤ n |] ==> (D n = b)" by (simp add: partition) lemma lemma_partition_lt_gen [rule_format]: "partition(a,b) D & m + Suc d ≤ n & n ≤ (psize D) --> D(m) < D(m + Suc d)" apply (induct "d", auto simp add: partition) apply (blast dest: Suc_le_lessD intro: less_le_trans order_less_trans) done lemma less_eq_add_Suc: "m < n ==> ∃d. n = m + Suc d" by (auto simp add: less_iff_Suc_add) lemma partition_lt_gen: "[|partition(a,b) D; m < n; n ≤ (psize D)|] ==> D(m) < D(n)" by (auto dest: less_eq_add_Suc intro: lemma_partition_lt_gen) lemma partition_lt: "partition(a,b) D ==> n < (psize D) ==> D(0) < D(Suc n)" apply (induct "n") apply (auto simp add: partition) done lemma partition_le: "partition(a,b) D ==> a ≤ b" apply (frule partition [THEN iffD1], safe) apply (drule_tac x = "psize D" and P="%n. psize D ≤ n --> ?P n" in spec, safe) apply (case_tac "psize D = 0") apply (drule_tac [2] n = "psize D - 1" in partition_lt, auto) done lemma partition_gt: "[|partition(a,b) D; n < (psize D)|] ==> D(n) < D(psize D)" by (auto intro: partition_lt_gen) lemma partition_eq: "partition(a,b) D ==> ((a = b) = (psize D = 0))" apply (frule partition [THEN iffD1], safe) apply (rotate_tac 2) apply (drule_tac x = "psize D" in spec) apply (rule ccontr) apply (drule_tac n = "psize D - 1" in partition_lt) apply auto done lemma partition_lb: "partition(a,b) D ==> a ≤ D(r)" apply (frule partition [THEN iffD1], safe) apply (induct "r") apply (cut_tac [2] y = "Suc r" and x = "psize D" in linorder_le_less_linear) apply (auto intro: partition_le) apply (drule_tac x = r in spec) apply arith; done lemma partition_lb_lt: "[| partition(a,b) D; psize D ~= 0 |] ==> a < D(Suc n)" apply (rule_tac t = a in partition_lhs [THEN subst], assumption) apply (cut_tac x = "Suc n" and y = "psize D" in linorder_le_less_linear) apply (frule partition [THEN iffD1], safe) apply (blast intro: partition_lt less_le_trans) apply (rotate_tac 3) apply (drule_tac x = "Suc n" in spec) apply (erule impE) apply (erule less_imp_le) apply (frule partition_rhs) apply (drule partition_gt[of _ _ _ 0], arith) apply (simp (no_asm_simp)) done lemma partition_ub: "partition(a,b) D ==> D(r) ≤ b" apply (frule partition [THEN iffD1]) apply (cut_tac x = "psize D" and y = r in linorder_le_less_linear, safe, blast) apply (subgoal_tac "∀x. D ((psize D) - x) ≤ b") apply (rotate_tac 4) apply (drule_tac x = "psize D - r" in spec) apply (subgoal_tac "psize D - (psize D - r) = r") apply simp apply arith apply safe apply (induct_tac "x") apply (simp (no_asm), blast) apply (case_tac "psize D - Suc n = 0") apply (erule_tac V = "∀n. psize D ≤ n --> D n = b" in thin_rl) apply (simp (no_asm_simp) add: partition_le) apply (rule order_trans) prefer 2 apply assumption apply (subgoal_tac "psize D - n = Suc (psize D - Suc n)") prefer 2 apply arith apply (drule_tac x = "psize D - Suc n" in spec, simp) done lemma partition_ub_lt: "[| partition(a,b) D; n < psize D |] ==> D(n) < b" by (blast intro: partition_rhs [THEN subst] partition_gt) lemma lemma_partition_append1: "[| partition (a, b) D1; partition (b, c) D2 |] ==> (∀n < psize D1 + psize D2. (if n < psize D1 then D1 n else D2 (n - psize D1)) < (if Suc n < psize D1 then D1 (Suc n) else D2 (Suc n - psize D1))) & (∀n ≥ psize D1 + psize D2. (if n < psize D1 then D1 n else D2 (n - psize D1)) = (if psize D1 + psize D2 < psize D1 then D1 (psize D1 + psize D2) else D2 (psize D1 + psize D2 - psize D1)))" apply (auto intro: partition_lt_gen) apply (subgoal_tac "psize D1 = Suc n") apply (auto intro!: partition_lt_gen simp add: partition_lhs partition_ub_lt) apply (auto intro!: partition_rhs2 simp add: partition_rhs split: nat_diff_split) done lemma lemma_psize1: "[| partition (a, b) D1; partition (b, c) D2; N < psize D1 |] ==> D1(N) < D2 (psize D2)" apply (rule_tac y = "D1 (psize D1)" in order_less_le_trans) apply (erule partition_gt) apply (auto simp add: partition_rhs partition_le) done lemma lemma_partition_append2: "[| partition (a, b) D1; partition (b, c) D2 |] ==> psize (%n. if n < psize D1 then D1 n else D2 (n - psize D1)) = psize D1 + psize D2" apply (unfold psize_def [of "%n. if n < psize D1 then D1 n else D2 (n - psize D1)"]) apply (rule some1_equality) prefer 2 apply (blast intro!: lemma_partition_append1) apply (rule ex1I, rule lemma_partition_append1) apply (simp_all split: split_if_asm) txt{*The case @{term "N < psize D1"}*} apply (drule_tac x = "psize D1 + psize D2" and P="%n. ?P n & ?Q n" in spec) apply (force dest: lemma_psize1) apply (rule order_antisym); txt{*The case @{term "psize D1 ≤ N"}: proving @{term "N ≤ psize D1 + psize D2"}*} apply (drule_tac x = "psize D1 + psize D2" in spec) apply (simp add: partition_rhs2) apply (case_tac "N - psize D1 < psize D2") prefer 2 apply arith txt{*Proving @{term "psize D1 + psize D2 ≤ N"}*} apply (drule_tac x = "psize D1 + psize D2" and P="%n. N≤n --> ?P n" in spec, simp) apply (drule_tac a = b and b = c in partition_gt, assumption, simp) done lemma tpart_eq_lhs_rhs: "[|psize D = 0; tpart(a,b) (D,p)|] ==> a = b" by (auto simp add: tpart_def partition_eq) lemma tpart_partition: "tpart(a,b) (D,p) ==> partition(a,b) D" by (simp add: tpart_def) lemma partition_append: "[| tpart(a,b) (D1,p1); fine(g) (D1,p1); tpart(b,c) (D2,p2); fine(g) (D2,p2) |] ==> ∃D p. tpart(a,c) (D,p) & fine(g) (D,p)" apply (rule_tac x = "%n. if n < psize D1 then D1 n else D2 (n - psize D1)" in exI) apply (rule_tac x = "%n. if n < psize D1 then p1 n else p2 (n - psize D1)" in exI) apply (case_tac "psize D1 = 0") apply (auto dest: tpart_eq_lhs_rhs) prefer 2 apply (simp add: fine_def lemma_partition_append2 [OF tpart_partition tpart_partition]) --{*But must not expand @{term fine} in other subgoals*} apply auto apply (subgoal_tac "psize D1 = Suc n") prefer 2 apply arith apply (drule tpart_partition [THEN partition_rhs]) apply (drule tpart_partition [THEN partition_lhs]) apply (auto split: nat_diff_split) apply (auto simp add: tpart_def) defer 1 apply (subgoal_tac "psize D1 = Suc n") prefer 2 apply arith apply (drule partition_rhs) apply (drule partition_lhs, auto) apply (simp split: nat_diff_split) apply (subst partition) apply (subst (1 2) lemma_partition_append2, assumption+) apply (rule conjI) apply (simp add: partition_lhs) apply (drule lemma_partition_append1) apply assumption; apply (simp add: partition_rhs) done text{*We can always find a division that is fine wrt any gauge*} lemma partition_exists: "[| a ≤ b; gauge(%x. a ≤ x & x ≤ b) g |] ==> ∃D p. tpart(a,b) (D,p) & fine g (D,p)" apply (cut_tac P = "%(u,v). a ≤ u & v ≤ b --> (∃D p. tpart (u,v) (D,p) & fine (g) (D,p))" in lemma_BOLZANO2) apply safe apply (blast intro: order_trans)+ apply (auto intro: partition_append) apply (case_tac "a ≤ x & x ≤ b") apply (rule_tac [2] x = 1 in exI, auto) apply (rule_tac x = "g x" in exI) apply (auto simp add: gauge_def) apply (rule_tac x = "%n. if n = 0 then aa else ba" in exI) apply (rule_tac x = "%n. if n = 0 then x else ba" in exI) apply (auto simp add: tpart_def fine_def) done text{*Lemmas about combining gauges*} lemma gauge_min: "[| gauge(E) g1; gauge(E) g2 |] ==> gauge(E) (%x. if g1(x) < g2(x) then g1(x) else g2(x))" by (simp add: gauge_def) lemma fine_min: "fine (%x. if g1(x) < g2(x) then g1(x) else g2(x)) (D,p) ==> fine(g1) (D,p) & fine(g2) (D,p)" by (auto simp add: fine_def split: split_if_asm) text{*The integral is unique if it exists*} lemma Integral_unique: "[| a ≤ b; Integral(a,b) f k1; Integral(a,b) f k2 |] ==> k1 = k2" apply (simp add: Integral_def) apply (drule_tac x = "¦k1 - k2¦ /2" in spec)+ apply auto apply (drule gauge_min, assumption) apply (drule_tac g = "%x. if g x < ga x then g x else ga x" in partition_exists, assumption, auto) apply (drule fine_min) apply (drule spec)+ apply auto apply (subgoal_tac "¦(rsum (D,p) f - k2) - (rsum (D,p) f - k1)¦ < ¦k1 - k2¦") apply arith apply (drule add_strict_mono, assumption) apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] mult_less_cancel_right) done lemma Integral_zero [simp]: "Integral(a,a) f 0" apply (auto simp add: Integral_def) apply (rule_tac x = "%x. 1" in exI) apply (auto dest: partition_eq simp add: gauge_def tpart_def rsum_def) done lemma sumr_partition_eq_diff_bounds [simp]: "(∑n=0..<m. D (Suc n) - D n::real) = D(m) - D 0" by (induct "m", auto) lemma Integral_eq_diff_bounds: "a ≤ b ==> Integral(a,b) (%x. 1) (b - a)" apply (auto simp add: order_le_less rsum_def Integral_def) apply (rule_tac x = "%x. b - a" in exI) apply (auto simp add: gauge_def abs_less_iff tpart_def partition) done lemma Integral_mult_const: "a ≤ b ==> Integral(a,b) (%x. c) (c*(b - a))" apply (auto simp add: order_le_less rsum_def Integral_def) apply (rule_tac x = "%x. b - a" in exI) apply (auto simp add: setsum_right_distrib [symmetric] gauge_def abs_less_iff right_diff_distrib [symmetric] partition tpart_def) done lemma Integral_mult: "[| a ≤ b; Integral(a,b) f k |] ==> Integral(a,b) (%x. c * f x) (c * k)" apply (auto simp add: order_le_less dest: Integral_unique [OF order_refl Integral_zero]) apply (auto simp add: rsum_def Integral_def setsum_right_distrib[symmetric] mult_assoc) apply (rule_tac a2 = c in abs_ge_zero [THEN order_le_imp_less_or_eq, THEN disjE]) prefer 2 apply force apply (drule_tac x = "e/abs c" in spec, auto) apply (simp add: zero_less_mult_iff divide_inverse) apply (rule exI, auto) apply (drule spec)+ apply auto apply (rule_tac z1 = "inverse (abs c)" in real_mult_less_iff1 [THEN iffD1]) apply (auto simp add: abs_mult divide_inverse [symmetric] right_diff_distrib [symmetric]) done text{*Fundamental theorem of calculus (Part I)*} text{*"Straddle Lemma" : Swartz and Thompson: AMM 95(7) 1988 *} lemma choiceP: "∀x. P(x) --> (∃y. Q x y) ==> ∃f. (∀x. P(x) --> Q x (f x))" by (insert bchoice [of "Collect P" Q], simp) (*UNUSED lemma choice2: "∀x. (∃y. R(y) & (∃z. Q x y z)) ==> ∃f fa. (∀x. R(f x) & Q x (f x) (fa x))" *) (* new simplifications e.g. (y < x/n) = (y * n < x) are a real nuisance they break the original proofs and make new proofs longer!*) lemma strad1: "[|∀xa::real. xa ≠ x ∧ ¦xa - x¦ < s --> ¦(f xa - f x) / (xa - x) - f' x¦ * 2 < e; 0 < e; a ≤ x; x ≤ b; 0 < s|] ==> ∀z. ¦z - x¦ < s -->¦f z - f x - f' x * (z - x)¦ * 2 ≤ e * ¦z - x¦" apply auto apply (case_tac "0 < ¦z - x¦") prefer 2 apply (simp add: zero_less_abs_iff) apply (drule_tac x = z in spec) apply (rule_tac z1 = "¦inverse (z - x)¦" in real_mult_le_cancel_iff2 [THEN iffD1]) apply simp apply (simp del: abs_inverse abs_mult add: abs_mult [symmetric] mult_assoc [symmetric]) apply (subgoal_tac "inverse (z - x) * (f z - f x - f' x * (z - x)) = (f z - f x) / (z - x) - f' x") apply (simp add: abs_mult [symmetric] mult_ac diff_minus) apply (subst mult_commute) apply (simp add: left_distrib diff_minus) apply (simp add: mult_assoc divide_inverse) apply (simp add: left_distrib) done lemma lemma_straddle: "[| ∀x. a ≤ x & x ≤ b --> DERIV f x :> f'(x); 0 < e |] ==> ∃g. gauge(%x. a ≤ x & x ≤ b) g & (∀x u v. a ≤ u & u ≤ x & x ≤ v & v ≤ b & (v - u) < g(x) --> ¦(f(v) - f(u)) - (f'(x) * (v - u))¦ ≤ e * (v - u))" apply (simp add: gauge_def) apply (subgoal_tac "∀x. a ≤ x & x ≤ b --> (∃d > 0. ∀u v. u ≤ x & x ≤ v & (v - u) < d --> ¦(f (v) - f (u)) - (f' (x) * (v - u))¦ ≤ e * (v - u))") apply (drule choiceP, auto) apply (drule spec, auto) apply (auto simp add: DERIV_iff2 LIM_def) apply (drule_tac x = "e/2" in spec, auto) apply (frule strad1, assumption+) apply (rule_tac x = s in exI, auto) apply (rule_tac x = u and y = v in linorder_cases, auto) apply (rule_tac y = "¦(f (v) - f (x)) - (f' (x) * (v - x))¦ + ¦(f (x) - f (u)) - (f' (x) * (x - u))¦" in order_trans) apply (rule abs_triangle_ineq [THEN [2] order_trans]) apply (simp add: right_diff_distrib) apply (rule_tac t = "e* (v - u)" in real_sum_of_halves [THEN subst]) apply (rule add_mono) apply (rule_tac y = "(e/2) * ¦v - x¦" in order_trans) prefer 2 apply simp apply (erule_tac [!] V= "∀x'. x' ~= x & ¦x' - x¦ < s --> ?P x'" in thin_rl) apply (drule_tac x = v in spec, simp add: times_divide_eq) apply (drule_tac x = u in spec, auto) apply (subgoal_tac "¦f u - f x - f' x * (u - x)¦ = ¦f x - f u - f' x * (x - u)¦") apply (rule order_trans) apply (auto simp add: abs_le_iff) apply (simp add: right_diff_distrib) done lemma FTC1: "[|a ≤ b; ∀x. a ≤ x & x ≤ b --> DERIV f x :> f'(x) |] ==> Integral(a,b) f' (f(b) - f(a))" apply (drule order_le_imp_less_or_eq, auto) apply (auto simp add: Integral_def) apply (rule ccontr) apply (subgoal_tac "∀e > 0. ∃g. gauge (%x. a ≤ x & x ≤ b) g & (∀D p. tpart (a, b) (D, p) & fine g (D, p) --> ¦rsum (D, p) f' - (f b - f a)¦ ≤ e)") apply (rotate_tac 3) apply (drule_tac x = "e/2" in spec, auto) apply (drule spec, auto) apply ((drule spec)+, auto) apply (drule_tac e = "ea/ (b - a)" in lemma_straddle) apply (auto simp add: zero_less_divide_iff) apply (rule exI) apply (auto simp add: tpart_def rsum_def) apply (subgoal_tac "(∑n=0..<psize D. f(D(Suc n)) - f(D n)) = f b - f a") prefer 2 apply (cut_tac D = "%n. f (D n)" and m = "psize D" in sumr_partition_eq_diff_bounds) apply (simp add: partition_lhs partition_rhs) apply (drule sym, simp) apply (simp (no_asm) add: setsum_subtractf[symmetric]) apply (rule setsum_abs [THEN order_trans]) apply (subgoal_tac "ea = (∑n=0..<psize D. (ea / (b - a)) * (D (Suc n) - (D n)))") apply (simp add: abs_minus_commute) apply (rule_tac t = ea in ssubst, assumption) apply (rule setsum_mono) apply (rule_tac [2] setsum_right_distrib [THEN subst]) apply (auto simp add: partition_rhs partition_lhs partition_lb partition_ub fine_def) done lemma Integral_subst: "[| Integral(a,b) f k1; k2=k1 |] ==> Integral(a,b) f k2" by simp lemma Integral_add: "[| a ≤ b; b ≤ c; Integral(a,b) f' k1; Integral(b,c) f' k2; ∀x. a ≤ x & x ≤ c --> DERIV f x :> f' x |] ==> Integral(a,c) f' (k1 + k2)" apply (rule FTC1 [THEN Integral_subst], auto) apply (frule FTC1, auto) apply (frule_tac a = b in FTC1, auto) apply (drule_tac x = x in spec, auto) apply (drule_tac ?k2.0 = "f b - f a" in Integral_unique) apply (drule_tac [3] ?k2.0 = "f c - f b" in Integral_unique, auto) done lemma partition_psize_Least: "partition(a,b) D ==> psize D = (LEAST n. D(n) = b)" apply (auto intro!: Least_equality [symmetric] partition_rhs) apply (auto dest: partition_ub_lt simp add: linorder_not_less [symmetric]) done lemma lemma_partition_bounded: "partition (a, c) D ==> ~ (∃n. c < D(n))" apply safe apply (drule_tac r = n in partition_ub, auto) done lemma lemma_partition_eq: "partition (a, c) D ==> D = (%n. if D n < c then D n else c)" apply (rule ext, auto) apply (auto dest!: lemma_partition_bounded) apply (drule_tac x = n in spec, auto) done lemma lemma_partition_eq2: "partition (a, c) D ==> D = (%n. if D n ≤ c then D n else c)" apply (rule ext, auto) apply (auto dest!: lemma_partition_bounded) apply (drule_tac x = n in spec, auto) done lemma partition_lt_Suc: "[| partition(a,b) D; n < psize D |] ==> D n < D (Suc n)" by (auto simp add: partition) lemma tpart_tag_eq: "tpart(a,c) (D,p) ==> p = (%n. if D n < c then p n else c)" apply (rule ext) apply (auto simp add: tpart_def) apply (drule linorder_not_less [THEN iffD1]) apply (drule_tac r = "Suc n" in partition_ub) apply (drule_tac x = n in spec, auto) done subsection{*Lemmas for Additivity Theorem of Gauge Integral*} lemma lemma_additivity1: "[| a ≤ D n; D n < b; partition(a,b) D |] ==> n < psize D" by (auto simp add: partition linorder_not_less [symmetric]) lemma lemma_additivity2: "[| a ≤ D n; partition(a,D n) D |] ==> psize D ≤ n" apply (rule ccontr, drule not_leE) apply (frule partition [THEN iffD1], safe) apply (frule_tac r = "Suc n" in partition_ub) apply (auto dest!: spec) done lemma partition_eq_bound: "[| partition(a,b) D; psize D < m |] ==> D(m) = D(psize D)" by (auto simp add: partition) lemma partition_ub2: "[| partition(a,b) D; psize D < m |] ==> D(r) ≤ D(m)" by (simp add: partition partition_ub) lemma tag_point_eq_partition_point: "[| tpart(a,b) (D,p); psize D ≤ m |] ==> p(m) = D(m)" apply (simp add: tpart_def, auto) apply (drule_tac x = m in spec) apply (auto simp add: partition_rhs2) done lemma partition_lt_cancel: "[| partition(a,b) D; D m < D n |] ==> m < n" apply (cut_tac less_linear [of n "psize D"], auto) apply (cut_tac less_linear [of m n]) apply (cut_tac less_linear [of m "psize D"]) apply (auto dest: partition_gt) apply (drule_tac n = m in partition_lt_gen, auto) apply (frule partition_eq_bound) apply (drule_tac [2] partition_gt, auto) apply (metis dense_linear_order_class.dlo_simps(8) le_def partition_rhs partition_rhs2) apply (metis Nat.le_less_trans dense_linear_order_class.dlo_simps(8) nat_le_linear partition_eq_bound partition_rhs2) done lemma lemma_additivity4_psize_eq: "[| a ≤ D n; D n < b; partition (a, b) D |] ==> psize (%x. if D x < D n then D(x) else D n) = n" apply (unfold psize_def) apply (frule lemma_additivity1) apply (assumption, assumption) apply (rule some_equality) apply (auto intro: partition_lt_Suc) apply (drule_tac n = n in partition_lt_gen, assumption) apply (arith, arith) apply (cut_tac m = na and n = "psize D" in Nat.less_linear) apply (auto dest: partition_lt_cancel) apply (rule_tac x=N and y=n in linorder_cases) apply (drule_tac x = n and P="%m. N ≤ m --> ?f m = ?g m" in spec, simp) apply (drule_tac n = n in partition_lt_gen, auto) apply (drule_tac x = n in spec) apply (simp split: split_if_asm) done lemma lemma_psize_left_less_psize: "partition (a, b) D ==> psize (%x. if D x < D n then D(x) else D n) ≤ psize D" apply (frule_tac r = n in partition_ub) apply (drule_tac x = "D n" in order_le_imp_less_or_eq) apply (auto simp add: lemma_partition_eq [symmetric]) apply (frule_tac r = n in partition_lb) apply (drule (2) lemma_additivity4_psize_eq) apply (rule ccontr, auto) apply (frule_tac not_leE [THEN [2] partition_eq_bound]) apply (auto simp add: partition_rhs) done lemma lemma_psize_left_less_psize2: "[| partition(a,b) D; na < psize (%x. if D x < D n then D(x) else D n) |] ==> na < psize D" by (erule lemma_psize_left_less_psize [THEN [2] less_le_trans]) lemma lemma_additivity3: "[| partition(a,b) D; D na < D n; D n < D (Suc na); n < psize D |] ==> False" by (metis not_less_eq partition_lt_cancel real_of_nat_less_iff) lemma psize_const [simp]: "psize (%x. k) = 0" by (auto simp add: psize_def) lemma lemma_additivity3a: "[| partition(a,b) D; D na < D n; D n < D (Suc na); na < psize D |] ==> False" apply (frule_tac m = n in partition_lt_cancel) apply (auto intro: lemma_additivity3) done lemma better_lemma_psize_right_eq1: "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) ≤ psize D - n" apply (simp add: psize_def [of "(%x. D (x + n))"]); apply (rule_tac a = "psize D - n" in someI2, auto) apply (simp add: partition less_diff_conv) apply (simp add: le_diff_conv partition_rhs2 split: nat_diff_split) apply (drule_tac x = "psize D - n" in spec, auto) apply (frule partition_rhs, safe) apply (frule partition_lt_cancel, assumption) apply (drule partition [THEN iffD1], safe) apply (subgoal_tac "~ D (psize D - n + n) < D (Suc (psize D - n + n))") apply blast apply (drule_tac x = "Suc (psize D)" and P="%n. ?P n --> D n = D (psize D)" in spec) apply simp done lemma psize_le_n: "partition (a, D n) D ==> psize D ≤ n" apply (rule ccontr, drule not_leE) apply (frule partition_lt_Suc, assumption) apply (frule_tac r = "Suc n" in partition_ub, auto) done lemma better_lemma_psize_right_eq1a: "partition(a,D n) D ==> psize (%x. D (x + n)) ≤ psize D - n" apply (simp add: psize_def [of "(%x. D (x + n))"]); apply (rule_tac a = "psize D - n" in someI2, auto) apply (simp add: partition less_diff_conv) apply (simp add: le_diff_conv) apply (case_tac "psize D ≤ n") apply (force intro: partition_rhs2) apply (simp add: partition linorder_not_le) apply (rule ccontr, drule not_leE) apply (frule psize_le_n) apply (drule_tac x = "psize D - n" in spec, simp) apply (drule partition [THEN iffD1], safe) apply (drule_tac x = "Suc n" and P="%na. ?s ≤ na --> D na = D n" in spec, auto) done lemma better_lemma_psize_right_eq: "partition(a,b) D ==> psize (%x. D (x + n)) ≤ psize D - n" apply (frule_tac r1 = n in partition_ub [THEN order_le_imp_less_or_eq]) apply (blast intro: better_lemma_psize_right_eq1a better_lemma_psize_right_eq1) done lemma lemma_psize_right_eq1: "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) ≤ psize D" apply (simp add: psize_def [of "(%x. D (x + n))"]) apply (rule_tac a = "psize D - n" in someI2, auto) apply (simp add: partition less_diff_conv) apply (subgoal_tac "n ≤ psize D") apply (simp add: partition le_diff_conv) apply (rule ccontr, drule not_leE) apply (drule_tac less_imp_le [THEN [2] partition_rhs2], assumption, simp) apply (drule_tac x = "psize D" in spec) apply (simp add: partition) done (* should be combined with previous theorem; also proof has redundancy *) lemma lemma_psize_right_eq1a: "partition(a,D n) D ==> psize (%x. D (x + n)) ≤ psize D" apply (simp add: psize_def [of "(%x. D (x + n))"]); apply (rule_tac a = "psize D - n" in someI2, auto) apply (simp add: partition less_diff_conv) apply (case_tac "psize D ≤ n") apply (force intro: partition_rhs2 simp add: le_diff_conv) apply (simp add: partition le_diff_conv) apply (rule ccontr, drule not_leE) apply (drule_tac x = "psize D" in spec) apply (simp add: partition) done lemma lemma_psize_right_eq: "[| partition(a,b) D |] ==> psize (%x. D (x + n)) ≤ psize D" apply (frule_tac r1 = n in partition_ub [THEN order_le_imp_less_or_eq]) apply (blast intro: lemma_psize_right_eq1a lemma_psize_right_eq1) done lemma tpart_left1: "[| a ≤ D n; tpart (a, b) (D, p) |] ==> tpart(a, D n) (%x. if D x < D n then D(x) else D n, %x. if D x < D n then p(x) else D n)" apply (frule_tac r = n in tpart_partition [THEN partition_ub]) apply (drule_tac x = "D n" in order_le_imp_less_or_eq) apply (auto simp add: tpart_partition [THEN lemma_partition_eq, symmetric] tpart_tag_eq [symmetric]) apply (frule_tac tpart_partition [THEN [3] lemma_additivity1]) apply (auto simp add: tpart_def) apply (drule_tac [2] linorder_not_less [THEN iffD1, THEN order_le_imp_less_or_eq], auto) prefer 3 apply (drule_tac x=na in spec, arith) prefer 2 apply (blast dest: lemma_additivity3) apply (frule (2) lemma_additivity4_psize_eq) apply (rule partition [THEN iffD2]) apply (frule partition [THEN iffD1]) apply safe apply (auto simp add: partition_lt_gen) apply (drule (1) partition_lt_cancel, arith) done lemma fine_left1: "[| a ≤ D n; tpart (a, b) (D, p); gauge (%x. a ≤ x & x ≤ D n) g; fine (%x. if x < D n then min (g x) ((D n - x)/ 2) else if x = D n then min (g (D n)) (ga (D n)) else min (ga x) ((x - D n)/ 2)) (D, p) |] ==> fine g (%x. if D x < D n then D(x) else D n, %x. if D x < D n then p(x) else D n)" apply (auto simp add: fine_def tpart_def gauge_def) apply (frule_tac [!] na=na in lemma_psize_left_less_psize2) apply (drule_tac [!] x = na in spec, auto) apply (drule_tac [!] x = na in spec, auto) apply (auto dest: lemma_additivity3a simp add: split_if_asm) done lemma tpart_right1: "[| a ≤ D n; tpart (a, b) (D, p) |] ==> tpart(D n, b) (%x. D(x + n),%x. p(x + n))" apply (simp add: tpart_def partition_def, safe) apply (rule_tac x = "N - n" in exI, auto) done lemma fine_right1: "[| a ≤ D n; tpart (a, b) (D, p); gauge (%x. D n ≤ x & x ≤ b) ga; fine (%x. if x < D n then min (g x) ((D n - x)/ 2) else if x = D n then min (g (D n)) (ga (D n)) else min (ga x) ((x - D n)/ 2)) (D, p) |] ==> fine ga (%x. D(x + n),%x. p(x + n))" apply (auto simp add: fine_def gauge_def) apply (drule_tac x = "na + n" in spec) apply (frule_tac n = n in tpart_partition [THEN better_lemma_psize_right_eq], auto) apply (simp add: tpart_def, safe) apply (subgoal_tac "D n ≤ p (na + n)") apply (drule_tac y = "p (na + n)" in order_le_imp_less_or_eq) apply safe apply (simp split: split_if_asm, simp) apply (drule less_le_trans, assumption) apply (rotate_tac 5) apply (drule_tac x = "na + n" in spec, safe) apply (rule_tac y="D (na + n)" in order_trans) apply (case_tac "na = 0", auto) apply (erule partition_lt_gen [THEN order_less_imp_le]) apply arith apply arith done lemma rsum_add: "rsum (D, p) (%x. f x + g x) = rsum (D, p) f + rsum(D, p) g" by (simp add: rsum_def setsum_addf left_distrib) text{* Bartle/Sherbert: Theorem 10.1.5 p. 278 *} lemma Integral_add_fun: "[| a ≤ b; Integral(a,b) f k1; Integral(a,b) g k2 |] ==> Integral(a,b) (%x. f x + g x) (k1 + k2)" apply (simp add: Integral_def, auto) apply ((drule_tac x = "e/2" in spec)+) apply auto apply (drule gauge_min, assumption) apply (rule_tac x = " (%x. if ga x < gaa x then ga x else gaa x)" in exI) apply auto apply (drule fine_min) apply ((drule spec)+, auto) apply (drule_tac a = "¦rsum (D, p) f - k1¦ * 2" and c = "¦rsum (D, p) g - k2¦ * 2" in add_strict_mono, assumption) apply (auto simp only: rsum_add left_distrib [symmetric] mult_2_right [symmetric] real_mult_less_iff1) done lemma partition_lt_gen2: "[| partition(a,b) D; r < psize D |] ==> 0 < D (Suc r) - D r" by (auto simp add: partition) lemma lemma_Integral_le: "[| ∀x. a ≤ x & x ≤ b --> f x ≤ g x; tpart(a,b) (D,p) |] ==> ∀n ≤ psize D. f (p n) ≤ g (p n)" apply (simp add: tpart_def) apply (auto, frule partition [THEN iffD1], auto) apply (drule_tac x = "p n" in spec, auto) apply (case_tac "n = 0", simp) apply (rule partition_lt_gen [THEN order_less_le_trans, THEN order_less_imp_le], auto) apply (drule le_imp_less_or_eq, auto) apply (drule_tac [2] x = "psize D" in spec, auto) apply (drule_tac r = "Suc n" in partition_ub) apply (drule_tac x = n in spec, auto) done lemma lemma_Integral_rsum_le: "[| ∀x. a ≤ x & x ≤ b --> f x ≤ g x; tpart(a,b) (D,p) |] ==> rsum(D,p) f ≤ rsum(D,p) g" apply (simp add: rsum_def) apply (auto intro!: setsum_mono dest: tpart_partition [THEN partition_lt_gen2] dest!: lemma_Integral_le) done lemma Integral_le: "[| a ≤ b; ∀x. a ≤ x & x ≤ b --> f(x) ≤ g(x); Integral(a,b) f k1; Integral(a,b) g k2 |] ==> k1 ≤ k2" apply (simp add: Integral_def) apply (rotate_tac 2) apply (drule_tac x = "¦k1 - k2¦ /2" in spec) apply (drule_tac x = "¦k1 - k2¦ /2" in spec, auto) apply (drule gauge_min, assumption) apply (drule_tac g = "%x. if ga x < gaa x then ga x else gaa x" in partition_exists, assumption, auto) apply (drule fine_min) apply (drule_tac x = D in spec, drule_tac x = D in spec) apply (drule_tac x = p in spec, drule_tac x = p in spec, auto) apply (frule lemma_Integral_rsum_le, assumption) apply (subgoal_tac "¦(rsum (D,p) f - k1) - (rsum (D,p) g - k2)¦ < ¦k1 - k2¦") apply arith apply (drule add_strict_mono, assumption) apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] real_mult_less_iff1) done lemma Integral_imp_Cauchy: "(∃k. Integral(a,b) f k) ==> (∀e > 0. ∃g. gauge (%x. a ≤ x & x ≤ b) g & (∀D1 D2 p1 p2. tpart(a,b) (D1, p1) & fine g (D1,p1) & tpart(a,b) (D2, p2) & fine g (D2,p2) --> ¦rsum(D1,p1) f - rsum(D2,p2) f¦ < e))" apply (simp add: Integral_def, auto) apply (drule_tac x = "e/2" in spec, auto) apply (rule exI, auto) apply (frule_tac x = D1 in spec) apply (frule_tac x = D2 in spec) apply ((drule spec)+, auto) apply (erule_tac V = "0 < e" in thin_rl) apply (drule add_strict_mono, assumption) apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] real_mult_less_iff1) done lemma Cauchy_iff2: "Cauchy X = (∀j. (∃M. ∀m ≥ M. ∀n ≥ M. ¦X m - X n¦ < inverse(real (Suc j))))" apply (simp add: Cauchy_def, auto) apply (drule reals_Archimedean, safe) apply (drule_tac x = n in spec, auto) apply (rule_tac x = M in exI, auto) apply (drule_tac x = m in spec, simp) apply (drule_tac x = na in spec, auto) done lemma partition_exists2: "[| a ≤ b; ∀n. gauge (%x. a ≤ x & x ≤ b) (fa n) |] ==> ∀n. ∃D p. tpart (a, b) (D, p) & fine (fa n) (D, p)" by (blast dest: partition_exists) lemma monotonic_anti_derivative: fixes f g :: "real => real" shows "[| a ≤ b; ∀c. a ≤ c & c ≤ b --> f' c ≤ g' c; ∀x. DERIV f x :> f' x; ∀x. DERIV g x :> g' x |] ==> f b - f a ≤ g b - g a" apply (rule Integral_le, assumption) apply (auto intro: FTC1) done end
lemma partition_zero:
a = b ==> psize (λn. if n = 0 then a else b) = 0
lemma partition_one:
a < b ==> psize (λn. if n = 0 then a else b) = 1
lemma partition_single:
a ≤ b ==> Integration.partition (a, b) (λn. if n = 0 then a else b)
lemma partition_lhs:
Integration.partition (a, b) D ==> D 0 = a
lemma partition:
Integration.partition (a, b) D =
(D 0 = a ∧ (∀n<psize D. D n < D (Suc n)) ∧ (∀n≥psize D. D n = b))
lemma partition_rhs:
Integration.partition (a, b) D ==> D (psize D) = b
lemma partition_rhs2:
[| Integration.partition (a, b) D; psize D ≤ n |] ==> D n = b
lemma lemma_partition_lt_gen:
Integration.partition (a, b) D ∧ m + Suc d ≤ n ∧ n ≤ psize D
==> D m < D (m + Suc d)
lemma less_eq_add_Suc:
m < n ==> ∃d. n = m + Suc d
lemma partition_lt_gen:
[| Integration.partition (a, b) D; m < n; n ≤ psize D |] ==> D m < D n
lemma partition_lt:
[| Integration.partition (a, b) D; n < psize D |] ==> D 0 < D (Suc n)
lemma partition_le:
Integration.partition (a, b) D ==> a ≤ b
lemma partition_gt:
[| Integration.partition (a, b) D; n < psize D |] ==> D n < D (psize D)
lemma partition_eq:
Integration.partition (a, b) D ==> (a = b) = (psize D = 0)
lemma partition_lb:
Integration.partition (a, b) D ==> a ≤ D r
lemma partition_lb_lt:
[| Integration.partition (a, b) D; psize D ≠ 0 |] ==> a < D (Suc n)
lemma partition_ub:
Integration.partition (a, b) D ==> D r ≤ b
lemma partition_ub_lt:
[| Integration.partition (a, b) D; n < psize D |] ==> D n < b
lemma lemma_partition_append1:
[| Integration.partition (a, b) D1.0; Integration.partition (b, c) D2.0 |]
==> (∀n<psize D1.0 + psize D2.0.
(if n < psize D1.0 then D1.0 n else D2.0 (n - psize D1.0))
< (if Suc n < psize D1.0 then D1.0 (Suc n)
else D2.0 (Suc n - psize D1.0))) ∧
(∀n≥psize D1.0 + psize D2.0.
(if n < psize D1.0 then D1.0 n else D2.0 (n - psize D1.0)) =
(if psize D1.0 + psize D2.0 < psize D1.0
then D1.0 (psize D1.0 + psize D2.0)
else D2.0 (psize D1.0 + psize D2.0 - psize D1.0)))
lemma lemma_psize1:
[| Integration.partition (a, b) D1.0; Integration.partition (b, c) D2.0;
N < psize D1.0 |]
==> D1.0 N < D2.0 (psize D2.0)
lemma lemma_partition_append2:
[| Integration.partition (a, b) D1.0; Integration.partition (b, c) D2.0 |]
==> psize (λn. if n < psize D1.0 then D1.0 n else D2.0 (n - psize D1.0)) =
psize D1.0 + psize D2.0
lemma tpart_eq_lhs_rhs:
[| psize D = 0; tpart (a, b) (D, p) |] ==> a = b
lemma tpart_partition:
tpart (a, b) (D, p) ==> Integration.partition (a, b) D
lemma partition_append:
[| tpart (a, b) (D1.0, p1.0); fine g (D1.0, p1.0); tpart (b, c) (D2.0, p2.0);
fine g (D2.0, p2.0) |]
==> ∃D p. tpart (a, c) (D, p) ∧ fine g (D, p)
lemma partition_exists:
[| a ≤ b; gauge (λx. a ≤ x ∧ x ≤ b) g |]
==> ∃D p. tpart (a, b) (D, p) ∧ fine g (D, p)
lemma gauge_min:
[| gauge E g1.0; gauge E g2.0 |]
==> gauge E (λx. if g1.0 x < g2.0 x then g1.0 x else g2.0 x)
lemma fine_min:
fine (λx. if g1.0 x < g2.0 x then g1.0 x else g2.0 x) (D, p)
==> fine g1.0 (D, p) ∧ fine g2.0 (D, p)
lemma Integral_unique:
[| a ≤ b; Integral (a, b) f k1.0; Integral (a, b) f k2.0 |] ==> k1.0 = k2.0
lemma Integral_zero:
Integral (a, a) f 0
lemma sumr_partition_eq_diff_bounds:
(∑n = 0..<m. D (Suc n) - D n) = D m - D 0
lemma Integral_eq_diff_bounds:
a ≤ b ==> Integral (a, b) (λx. 1) (b - a)
lemma Integral_mult_const:
a ≤ b ==> Integral (a, b) (λx. c) (c * (b - a))
lemma Integral_mult:
[| a ≤ b; Integral (a, b) f k |] ==> Integral (a, b) (λx. c * f x) (c * k)
lemma choiceP:
∀x. P x --> (∃y. Q x y) ==> ∃f. ∀x. P x --> Q x (f x)
lemma strad1:
[| ∀xa. xa ≠ x ∧ ¦xa - x¦ < s --> ¦(f xa - f x) / (xa - x) - f' x¦ * 2 < e;
0 < e; a ≤ x; x ≤ b; 0 < s |]
==> ∀z. ¦z - x¦ < s --> ¦f z - f x - f' x * (z - x)¦ * 2 ≤ e * ¦z - x¦
lemma lemma_straddle:
[| ∀x. a ≤ x ∧ x ≤ b --> DERIV f x :> f' x; 0 < e |]
==> ∃g. gauge (λx. a ≤ x ∧ x ≤ b) g ∧
(∀x u v.
a ≤ u ∧ u ≤ x ∧ x ≤ v ∧ v ≤ b ∧ v - u < g x -->
¦f v - f u - f' x * (v - u)¦ ≤ e * (v - u))
lemma FTC1:
[| a ≤ b; ∀x. a ≤ x ∧ x ≤ b --> DERIV f x :> f' x |]
==> Integral (a, b) f' (f b - f a)
lemma Integral_subst:
[| Integral (a, b) f k1.0; k2.0 = k1.0 |] ==> Integral (a, b) f k2.0
lemma Integral_add:
[| a ≤ b; b ≤ c; Integral (a, b) f' k1.0; Integral (b, c) f' k2.0;
∀x. a ≤ x ∧ x ≤ c --> DERIV f x :> f' x |]
==> Integral (a, c) f' (k1.0 + k2.0)
lemma partition_psize_Least:
Integration.partition (a, b) D ==> psize D = (LEAST n. D n = b)
lemma lemma_partition_bounded:
Integration.partition (a, c) D ==> ¬ (∃n. c < D n)
lemma lemma_partition_eq:
Integration.partition (a, c) D ==> D = (λn. if D n < c then D n else c)
lemma lemma_partition_eq2:
Integration.partition (a, c) D ==> D = (λn. if D n ≤ c then D n else c)
lemma partition_lt_Suc:
[| Integration.partition (a, b) D; n < psize D |] ==> D n < D (Suc n)
lemma tpart_tag_eq:
tpart (a, c) (D, p) ==> p = (λn. if D n < c then p n else c)
lemma lemma_additivity1:
[| a ≤ D n; D n < b; Integration.partition (a, b) D |] ==> n < psize D
lemma lemma_additivity2:
[| a ≤ D n; Integration.partition (a, D n) D |] ==> psize D ≤ n
lemma partition_eq_bound:
[| Integration.partition (a, b) D; psize D < m |] ==> D m = D (psize D)
lemma partition_ub2:
[| Integration.partition (a, b) D; psize D < m |] ==> D r ≤ D m
lemma tag_point_eq_partition_point:
[| tpart (a, b) (D, p); psize D ≤ m |] ==> p m = D m
lemma partition_lt_cancel:
[| Integration.partition (a, b) D; D m < D n |] ==> m < n
lemma lemma_additivity4_psize_eq:
[| a ≤ D n; D n < b; Integration.partition (a, b) D |]
==> psize (λx. if D x < D n then D x else D n) = n
lemma lemma_psize_left_less_psize:
Integration.partition (a, b) D
==> psize (λx. if D x < D n then D x else D n) ≤ psize D
lemma lemma_psize_left_less_psize2:
[| Integration.partition (a, b) D;
na < psize (λx. if D x < D n then D x else D n) |]
==> na < psize D
lemma lemma_additivity3:
[| Integration.partition (a, b) D; D na < D n; D n < D (Suc na); n < psize D |]
==> False
lemma psize_const:
psize (λx. k) = 0
lemma lemma_additivity3a:
[| Integration.partition (a, b) D; D na < D n; D n < D (Suc na); na < psize D |]
==> False
lemma better_lemma_psize_right_eq1:
[| Integration.partition (a, b) D; D n < b |]
==> psize (λx. D (x + n)) ≤ psize D - n
lemma psize_le_n:
Integration.partition (a, D n) D ==> psize D ≤ n
lemma better_lemma_psize_right_eq1a:
Integration.partition (a, D n) D ==> psize (λx. D (x + n)) ≤ psize D - n
lemma better_lemma_psize_right_eq:
Integration.partition (a, b) D ==> psize (λx. D (x + n)) ≤ psize D - n
lemma lemma_psize_right_eq1:
[| Integration.partition (a, b) D; D n < b |]
==> psize (λx. D (x + n)) ≤ psize D
lemma lemma_psize_right_eq1a:
Integration.partition (a, D n) D ==> psize (λx. D (x + n)) ≤ psize D
lemma lemma_psize_right_eq:
Integration.partition (a, b) D ==> psize (λx. D (x + n)) ≤ psize D
lemma tpart_left1:
[| a ≤ D n; tpart (a, b) (D, p) |]
==> tpart (a, D n)
(λx. if D x < D n then D x else D n, λx. if D x < D n then p x else D n)
lemma fine_left1:
[| a ≤ D n; tpart (a, b) (D, p); gauge (λx. a ≤ x ∧ x ≤ D n) g;
fine (λx. if x < D n then min (g x) ((D n - x) / 2)
else if x = D n then min (g (D n)) (ga (D n))
else min (ga x) ((x - D n) / 2))
(D, p) |]
==> fine g
(λx. if D x < D n then D x else D n, λx. if D x < D n then p x else D n)
lemma tpart_right1:
[| a ≤ D n; tpart (a, b) (D, p) |]
==> tpart (D n, b) (λx. D (x + n), λx. p (x + n))
lemma fine_right1:
[| a ≤ D n; tpart (a, b) (D, p); gauge (λx. D n ≤ x ∧ x ≤ b) ga;
fine (λx. if x < D n then min (g x) ((D n - x) / 2)
else if x = D n then min (g (D n)) (ga (D n))
else min (ga x) ((x - D n) / 2))
(D, p) |]
==> fine ga (λx. D (x + n), λx. p (x + n))
lemma rsum_add:
rsum (D, p) (λx. f x + g x) = rsum (D, p) f + rsum (D, p) g
lemma Integral_add_fun:
[| a ≤ b; Integral (a, b) f k1.0; Integral (a, b) g k2.0 |]
==> Integral (a, b) (λx. f x + g x) (k1.0 + k2.0)
lemma partition_lt_gen2:
[| Integration.partition (a, b) D; r < psize D |] ==> 0 < D (Suc r) - D r
lemma lemma_Integral_le:
[| ∀x. a ≤ x ∧ x ≤ b --> f x ≤ g x; tpart (a, b) (D, p) |]
==> ∀n≤psize D. f (p n) ≤ g (p n)
lemma lemma_Integral_rsum_le:
[| ∀x. a ≤ x ∧ x ≤ b --> f x ≤ g x; tpart (a, b) (D, p) |]
==> rsum (D, p) f ≤ rsum (D, p) g
lemma Integral_le:
[| a ≤ b; ∀x. a ≤ x ∧ x ≤ b --> f x ≤ g x; Integral (a, b) f k1.0;
Integral (a, b) g k2.0 |]
==> k1.0 ≤ k2.0
lemma Integral_imp_Cauchy:
∃k. Integral (a, b) f k
==> ∀e>0. ∃g. gauge (λx. a ≤ x ∧ x ≤ b) g ∧
(∀D1 D2 p1 p2.
tpart (a, b) (D1, p1) ∧
fine g (D1, p1) ∧ tpart (a, b) (D2, p2) ∧ fine g (D2, p2) -->
¦rsum (D1, p1) f - rsum (D2, p2) f¦ < e)
lemma Cauchy_iff2:
Cauchy X = (∀j. ∃M. ∀m≥M. ∀n≥M. ¦X m - X n¦ < inverse (real (Suc j)))
lemma partition_exists2:
[| a ≤ b; ∀n. gauge (λx. a ≤ x ∧ x ≤ b) (fa n) |]
==> ∀n. ∃D p. tpart (a, b) (D, p) ∧ fine (fa n) (D, p)
lemma monotonic_anti_derivative:
[| a ≤ b; ∀c. a ≤ c ∧ c ≤ b --> f' c ≤ g' c; ∀x. DERIV f x :> f' x;
∀x. DERIV g x :> g' x |]
==> f b - f a ≤ g b - g a