# HG changeset patch # User immler # Date 1395133977 -3600 # Node ID 0268784f60daa298e462cfd0b5b5a1e5d0680b6f # Parent 2666cd7d380cb872c2381650f622db574339f51f use cbox to relax class constraints diff -r 2666cd7d380c -r 0268784f60da src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Mar 17 21:56:32 2014 +0100 +++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Tue Mar 18 10:12:57 2014 +0100 @@ -4426,21 +4426,21 @@ done lemma in_interval_interval_bij: - fixes a b u v x :: "'a::ordered_euclidean_space" - assumes "x \ {a..b}" - and "{u..v} \ {}" - shows "interval_bij (a, b) (u, v) x \ {u..v}" - apply (simp only: interval_bij_def split_conv mem_interval inner_setsum_left_Basis cong: ball_cong) + fixes a b u v x :: "'a::euclidean_space" + assumes "x \ cbox a b" + and "cbox u v \ {}" + shows "interval_bij (a, b) (u, v) x \ cbox u v" + apply (simp only: interval_bij_def split_conv mem_box inner_setsum_left_Basis cong: ball_cong) apply safe proof - fix i :: 'a assume i: "i \ Basis" - have "{a..b} \ {}" + have "cbox a b \ {}" using assms by auto with i have *: "a\i \ b\i" "u\i \ v\i" - using assms(2) by (auto simp add: interval_eq_empty interval) + using assms(2) by (auto simp add: box_eq_empty) have x: "a\i\x\i" "x\i\b\i" - using assms(1)[unfolded mem_interval] using i by auto + using assms(1)[unfolded mem_box] using i by auto have "0 \ (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i)" using * x by (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg) then show "u \ i \ u \ i + (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i)" diff -r 2666cd7d380c -r 0268784f60da src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Mon Mar 17 21:56:32 2014 +0100 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Tue Mar 18 10:12:57 2014 +0100 @@ -99,6 +99,9 @@ definition "vec x = (\ i. x)" +lemma interval_cbox_cart: "{a::real^'n..b} = cbox a b" + by (auto simp add: less_eq_vec_def mem_box Basis_vec_def inner_axis) + text{* Also the scalar-vector multiplication. *} definition vector_scalar_mult:: "'a::times \ 'a ^ 'n \ 'a ^ 'n" (infixl "*s" 70) @@ -911,19 +914,19 @@ lemma interval_cart: fixes a :: "real^'n" shows "box a b = {x::real^'n. \i. a$i < x$i \ x$i < b$i}" - and "{a .. b} = {x::real^'n. \i. a$i \ x$i \ x$i \ b$i}" - by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def mem_interval Basis_vec_def inner_axis) + and "cbox a b = {x::real^'n. \i. a$i \ x$i \ x$i \ b$i}" + by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def mem_box Basis_vec_def inner_axis) lemma mem_interval_cart: fixes a :: "real^'n" shows "x \ box a b \ (\i. a$i < x$i \ x$i < b$i)" - and "x \ {a .. b} \ (\i. a$i \ x$i \ x$i \ b$i)" + and "x \ cbox a b \ (\i. a$i \ x$i \ x$i \ b$i)" using interval_cart[of a b] by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def) lemma interval_eq_empty_cart: fixes a :: "real^'n" shows "(box a b = {} \ (\i. b$i \ a$i))" (is ?th1) - and "({a .. b} = {} \ (\i. b$i < a$i))" (is ?th2) + and "(cbox a b = {} \ (\i. b$i < a$i))" (is ?th2) proof - { fix i x assume as:"b$i \ a$i" and x:"x\box a b" hence "a $ i < x $ i \ x $ i < b $ i" unfolding mem_interval_cart by auto @@ -940,7 +943,7 @@ hence "box a b \ {}" using mem_interval_cart(1)[of "?x" a b] by auto } ultimately show ?th1 by blast - { fix i x assume as:"b$i < a$i" and x:"x\{a .. b}" + { fix i x assume as:"b$i < a$i" and x:"x\cbox a b" hence "a $ i \ x $ i \ x $ i \ b $ i" unfolding mem_interval_cart by auto hence "a$i \ b$i" by auto hence False using as by auto } @@ -952,22 +955,22 @@ hence "a$i \ ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \ b$i" unfolding vector_smult_component and vector_add_component by auto } - hence "{a .. b} \ {}" using mem_interval_cart(2)[of "?x" a b] by auto } + hence "cbox a b \ {}" using mem_interval_cart(2)[of "?x" a b] by auto } ultimately show ?th2 by blast qed lemma interval_ne_empty_cart: fixes a :: "real^'n" - shows "{a .. b} \ {} \ (\i. a$i \ b$i)" + shows "cbox a b \ {} \ (\i. a$i \ b$i)" and "box a b \ {} \ (\i. a$i < b$i)" unfolding interval_eq_empty_cart[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *) lemma subset_interval_imp_cart: fixes a :: "real^'n" - shows "(\i. a$i \ c$i \ d$i \ b$i) \ {c .. d} \ {a .. b}" - and "(\i. a$i < c$i \ d$i < b$i) \ {c .. d} \ box a b" - and "(\i. a$i \ c$i \ d$i \ b$i) \ box c d \ {a .. b}" + shows "(\i. a$i \ c$i \ d$i \ b$i) \ cbox c d \ cbox a b" + and "(\i. a$i < c$i \ d$i < b$i) \ cbox c d \ box a b" + and "(\i. a$i \ c$i \ d$i \ b$i) \ box c d \ cbox a b" and "(\i. a$i \ c$i \ d$i \ b$i) \ box c d \ box a b" unfolding subset_eq[unfolded Ball_def] unfolding mem_interval_cart by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *) @@ -978,49 +981,28 @@ apply (auto simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff) done -lemma interval_open_subset_closed_cart: - fixes a :: "real^'n" - shows "box a b \ {a .. b}" -proof (simp add: subset_eq, rule) - fix x - assume x: "x \box a b" - { fix i - have "a $ i \ x $ i" - using x order_less_imp_le[of "a$i" "x$i"] - by(simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff mem_interval Basis_vec_def inner_axis) - } - moreover - { fix i - have "x $ i \ b $ i" - using x order_less_imp_le[of "x$i" "b$i"] - by(simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff mem_interval Basis_vec_def inner_axis) - } - ultimately - show "a \ x \ x \ b" - by(simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff) -qed - lemma subset_interval_cart: fixes a :: "real^'n" - shows "{c .. d} \ {a .. b} \ (\i. c$i \ d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th1) - and "{c .. d} \ box a b \ (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) - and "box c d \ {a .. b} \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) + shows "cbox c d \ cbox a b \ (\i. c$i \ d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th1) + and "cbox c d \ box a b \ (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) + and "box c d \ cbox a b \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) and "box c d \ box a b \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th4) - using subset_interval[of c d a b] by (simp_all add: Basis_vec_def inner_axis) + using subset_box[of c d a b] by (simp_all add: Basis_vec_def inner_axis) lemma disjoint_interval_cart: fixes a::"real^'n" - shows "{a .. b} \ {c .. d} = {} \ (\i. (b$i < a$i \ d$i < c$i \ b$i < c$i \ d$i < a$i))" (is ?th1) - and "{a .. b} \ box c d = {} \ (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) - and "box a b \ {c .. d} = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) + shows "cbox a b \ cbox c d = {} \ (\i. (b$i < a$i \ d$i < c$i \ b$i < c$i \ d$i < a$i))" (is ?th1) + and "cbox a b \ box c d = {} \ (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) + and "box a b \ cbox c d = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) and "box a b \ box c d = {} \ (\i. (b$i \ a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th4) using disjoint_interval[of a b c d] by (simp_all add: Basis_vec_def inner_axis) lemma inter_interval_cart: fixes a :: "real^'n" - shows "{a .. b} \ {c .. d} = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" - unfolding set_eq_iff and Int_iff and mem_interval_cart - by auto + shows "cbox a b \ cbox c d = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" + unfolding inter_interval + by (auto simp: mem_box less_eq_vec_def) + (auto simp: Basis_vec_def inner_axis) lemma closed_interval_left_cart: fixes b :: "real^'n" @@ -1158,17 +1140,17 @@ by (rule convex_box_cart) (simp add: atLeast_def[symmetric] convex_real_interval) lemma unit_interval_convex_hull_cart: - "{0::real^'n .. 1} = convex hull {x. \i. (x$i = 0) \ (x$i = 1)}" - unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] + "cbox (0::real^'n) 1 = convex hull {x. \i. (x$i = 0) \ (x$i = 1)}" + unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] box_real[symmetric] by (rule arg_cong[where f="\x. convex hull x"]) (simp add: Basis_vec_def inner_axis) lemma cube_convex_hull_cart: assumes "0 < d" obtains s::"(real^'n) set" - where "finite s" "{x - (\ i. d) .. x + (\ i. d)} = convex hull s" + where "finite s" "cbox (x - (\ i. d)) (x + (\ i. d)) = convex hull s" proof - from assms obtain s where "finite s" - and "{x - setsum (op *\<^sub>R d) Basis..x + setsum (op *\<^sub>R d) Basis} = convex hull s" + and "cbox (x - setsum (op *\<^sub>R d) Basis) (x + setsum (op *\<^sub>R d) Basis) = convex hull s" by (rule cube_convex_hull) with that[of s] show thesis by (simp add: const_vector_cart) @@ -1344,16 +1326,16 @@ using component_le_norm_cart[of _ k] unfolding real_norm_def by auto lemma integral_component_eq_cart[simp]: - fixes f :: "'n::ordered_euclidean_space \ real^'m" + fixes f :: "'n::euclidean_space \ real^'m" assumes "f integrable_on s" shows "integral s (\x. f x $ k) = integral s f $ k" using integral_linear[OF assms(1) bounded_linear_component_cart,unfolded o_def] . lemma interval_split_cart: "{a..b::real^'n} \ {x. x$k \ c} = {a .. (\ i. if i = k then min (b$k) c else b$i)}" - "{a..b} \ {x. x$k \ c} = {(\ i. if i = k then max (a$k) c else a$i) .. b}" + "cbox a b \ {x. x$k \ c} = {(\ i. if i = k then max (a$k) c else a$i) .. b}" apply (rule_tac[!] set_eqI) - unfolding Int_iff mem_interval_cart mem_Collect_eq + unfolding Int_iff mem_interval_cart mem_Collect_eq interval_cbox_cart unfolding vec_lambda_beta by auto diff -r 2666cd7d380c -r 0268784f60da src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Mar 17 21:56:32 2014 +0100 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Mar 18 10:12:57 2014 +0100 @@ -1336,14 +1336,14 @@ text {* Balls, being convex, are connected. *} -lemma convex_box: +lemma convex_prod: assumes "\i. i \ Basis \ convex {x. P i x}" shows "convex {x. \i\Basis. P i (x\i)}" using assms unfolding convex_def by (auto simp: inner_add_left) lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\i\Basis. 0 \ x\i)}" - by (rule convex_box) (simp add: atLeast_def[symmetric] convex_real_interval) + by (rule convex_prod) (simp add: atLeast_def[symmetric] convex_real_interval) lemma convex_local_global_minimum: fixes s :: "'a::real_normed_vector set" @@ -3361,21 +3361,18 @@ ultimately show ?thesis by auto qed -lemma box_real: "box a b = {a<.. {}" using assms unfolding set_eq_iff - by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) + by (auto intro!: exI[of _ "(a + b) / 2"] simp: box) then show ?thesis - using interior_rel_interior_gen[of "{a..b}", symmetric] - by (simp split: split_if_asm add: interior_closed_interval box_real) + using interior_rel_interior_gen[of "cbox a b", symmetric] + by (simp split: split_if_asm del: box_real add: box_real[symmetric] interior_cbox) qed lemma rel_interior_real_semiline: @@ -3825,12 +3822,12 @@ apply simp apply auto done - have "continuous_on ({0..1} \ s \ t) (\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" + have "continuous_on ?X (\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" unfolding continuous_on by (rule ballI) (intro tendsto_intros) then show ?thesis unfolding * apply (rule compact_continuous_image) - apply (intro compact_Times compact_interval assms) + apply (intro compact_Times compact_Icc assms) done qed @@ -3854,7 +3851,7 @@ have "continuous_on ?T ?f" unfolding split_def continuous_on by (intro ballI tendsto_intros) moreover have "compact ?T" - by (intro compact_Times compact_interval insert) + by (intro compact_Times compact_Icc insert) ultimately have "compact (?f ` ?T)" by (rule compact_continuous_image) also have "?f ` ?T = convex hull (insert x A)" @@ -5083,7 +5080,7 @@ apply (rule continuous_at_imp_continuous_on) apply rule apply (intro continuous_intros) - apply (rule compact_interval) + apply (rule compact_Icc) done moreover have "{y. \u\0. u \ b / norm x \ y = u *\<^sub>R x} \ s \ {}" apply(rule *[OF _ assms(2)]) @@ -5673,8 +5670,8 @@ shows "is_interval s \ connected s" using is_interval_convex convex_connected by auto -lemma convex_interval: "convex {a .. b}" "convex (box a (b::'a::ordered_euclidean_space))" - apply (rule_tac[!] is_interval_convex) +lemma convex_box: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" + apply (rule_tac[!] is_interval_convex)+ using is_interval_interval apply auto done @@ -5741,35 +5738,35 @@ lemma ivt_increasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" - and "continuous_on {a .. b} f" + and "continuous_on (cbox a b) f" and "(f a)\k \ y" "y \ (f b)\k" - shows "\x\{a..b}. (f x)\k = y" -proof - - have "f a \ f ` {a..b}" "f b \ f ` {a..b}" + shows "\x\cbox a b. (f x)\k = y" +proof - + have "f a \ f ` cbox a b" "f b \ f ` cbox a b" apply (rule_tac[!] imageI) using assms(1) apply auto done then show ?thesis - using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y] - using connected_continuous_image[OF assms(2) convex_connected[OF convex_real_interval(5)]] + using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y] + using connected_continuous_image[OF assms(2) convex_connected[OF convex_box(1)]] using assms - by (auto intro!: imageI) + by auto qed lemma ivt_increasing_component_1: fixes f :: "real \ 'a::euclidean_space" - shows "a \ b \ \x\{a .. b}. continuous (at x) f \ - f a\k \ y \ y \ f b\k \ \x\{a..b}. (f x)\k = y" + shows "a \ b \ \x\cbox a b. continuous (at x) f \ + f a\k \ y \ y \ f b\k \ \x\cbox a b. (f x)\k = y" by (rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on) lemma ivt_decreasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" - and "continuous_on {a .. b} f" + and "continuous_on (cbox a b) f" and "(f b)\k \ y" and "y \ (f a)\k" - shows "\x\{a..b}. (f x)\k = y" + shows "\x\cbox a b. (f x)\k = y" apply (subst neg_equal_iff_equal[symmetric]) using ivt_increasing_component_on_1[of a b "\x. - f x" k "- y"] using assms using continuous_on_minus @@ -5778,8 +5775,8 @@ lemma ivt_decreasing_component_1: fixes f :: "real \ 'a::euclidean_space" - shows "a \ b \ \x\{a .. b}. continuous (at x) f \ - f b\k \ y \ y \ f a\k \ \x\{a..b}. (f x)\k = y" + shows "a \ b \ \x\cbox a b. continuous (at x) f \ + f b\k \ y \ y \ f a\k \ \x\cbox a b. (f x)\k = y" by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) @@ -5886,32 +5883,33 @@ by (simp add: f.add f.scaleR linear_iff) qed -lemma convex_hull_eq_real_interval: +lemma convex_hull_eq_real_cbox: fixes x y :: real assumes "x \ y" - shows "convex hull {x, y} = {x..y}" + shows "convex hull {x, y} = cbox x y" proof (rule hull_unique) - show "{x, y} \ {x..y}" using `x \ y` by auto - show "convex {x..y}" by (rule convex_interval) + show "{x, y} \ cbox x y" using `x \ y` by auto + show "convex (cbox x y)" + by (rule convex_box) next fix s assume "{x, y} \ s" and "convex s" - then show "{x..y} \ s" + then show "cbox x y \ s" unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def by - (clarify, simp (no_asm_use), fast) qed lemma unit_interval_convex_hull: defines "One \ \Basis" - shows "{0::'a::ordered_euclidean_space .. One} = convex hull {x. \i\Basis. (x\i = 0) \ (x\i = 1)}" + shows "cbox (0::'a::euclidean_space) One = convex hull {x. \i\Basis. (x\i = 0) \ (x\i = 1)}" (is "?int = convex hull ?points") proof - have One[simp]: "\i. i \ Basis \ One \ i = 1" by (simp add: One_def inner_setsum_left setsum_cases inner_Basis) - have "?int = {x. \i\Basis. x \ i \ {0..1}}" - by (auto simp add: eucl_le [where 'a='a]) - also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` {0..1})" + have "?int = {x. \i\Basis. x \ i \ cbox 0 1}" + by (auto simp: cbox_def) + also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` cbox 0 1)" by (simp only: box_eq_set_setsum_Basis) also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` (convex hull {0, 1}))" - by (simp only: convex_hull_eq_real_interval zero_le_one) + by (simp only: convex_hull_eq_real_cbox zero_le_one) also have "\ = (\i\Basis. convex hull ((\x. x *\<^sub>R i) ` {0, 1}))" by (simp only: convex_hull_linear_image linear_scaleR_left) also have "\ = convex hull (\i\Basis. (\x. x *\<^sub>R i) ` {0, 1})" @@ -5926,8 +5924,8 @@ text {* And this is a finite set of vertices. *} lemma unit_cube_convex_hull: - obtains s :: "'a::ordered_euclidean_space set" - where "finite s" and "{0 .. \Basis} = convex hull s" + obtains s :: "'a::euclidean_space set" + where "finite s" and "cbox 0 (\Basis) = convex hull s" apply (rule that[of "{x::'a. \i\Basis. x\i=0 \ x\i=1}"]) apply (rule finite_subset[of _ "(\s. (\i\Basis. (if i\s then 1 else 0) *\<^sub>R i)::'a) ` Pow Basis"]) prefer 3 @@ -5949,23 +5947,23 @@ lemma cube_convex_hull: assumes "d > 0" - obtains s :: "'a::ordered_euclidean_space set" where - "finite s" and "{x - (\i\Basis. d*\<^sub>Ri) .. x + (\i\Basis. d*\<^sub>Ri)} = convex hull s" + obtains s :: "'a::euclidean_space set" where + "finite s" and "cbox (x - (\i\Basis. d*\<^sub>Ri)) (x + (\i\Basis. d*\<^sub>Ri)) = convex hull s" proof - let ?d = "(\i\Basis. d*\<^sub>Ri)::'a" - have *: "{x - ?d .. x + ?d} = (\y. x - ?d + (2 * d) *\<^sub>R y) ` {0 .. \Basis}" + have *: "cbox (x - ?d) (x + ?d) = (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\Basis)" apply (rule set_eqI, rule) unfolding image_iff defer apply (erule bexE) proof - fix y - assume as: "y\{x - ?d .. x + ?d}" + assume as: "y\cbox (x - ?d) (x + ?d)" { fix i :: 'a assume i: "i \ Basis" have "x \ i \ d + y \ i" "y \ i \ d + x \ i" - using as[unfolded mem_interval, THEN bspec[where x=i]] i + using as[unfolded mem_box, THEN bspec[where x=i]] i by (auto simp: inner_simps) then have "1 \ inverse d * (x \ i - y \ i)" "1 \ inverse d * (y \ i - x \ i)" apply (rule_tac[!] mult_left_le_imp_le[OF _ assms]) @@ -5975,10 +5973,10 @@ then have "inverse d * (x \ i * 2) \ 2 + inverse d * (y \ i * 2)" "inverse d * (y \ i * 2) \ 2 + inverse d * (x \ i * 2)" by (auto simp add:field_simps) } - then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ {0..\Basis}" - unfolding mem_interval using assms + then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ cbox 0 (\Basis)" + unfolding mem_box using assms by (auto simp add: field_simps inner_simps) - then show "\z\{0..\Basis}. y = x - ?d + (2 * d) *\<^sub>R z" + then show "\z\cbox 0 (\Basis). y = x - ?d + (2 * d) *\<^sub>R z" apply - apply (rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI) using assms @@ -5986,9 +5984,9 @@ done next fix y z - assume as: "z\{0..\Basis}" "y = x - ?d + (2*d) *\<^sub>R z" + assume as: "z\cbox 0 (\Basis)" "y = x - ?d + (2*d) *\<^sub>R z" have "\i. i\Basis \ 0 \ d * (z \ i) \ d * (z \ i) \ d" - using assms as(1)[unfolded mem_interval] + using assms as(1)[unfolded mem_box] apply (erule_tac x=i in ballE) apply rule apply (rule mult_nonneg_nonneg) @@ -5997,17 +5995,17 @@ using assms apply auto done - then show "y \ {x - ?d..x + ?d}" - unfolding as(2) mem_interval + then show "y \ cbox (x - ?d) (x + ?d)" + unfolding as(2) mem_box apply - apply rule - using as(1)[unfolded mem_interval] + using as(1)[unfolded mem_box] apply (erule_tac x=i in ballE) using assms apply (auto simp: inner_simps) done qed - obtain s where "finite s" "{0::'a..\Basis} = convex hull s" + obtain s where "finite s" "cbox 0 (\Basis::'a) = convex hull s" using unit_cube_convex_hull by auto then show ?thesis apply (rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` s"]) @@ -6198,14 +6196,14 @@ def c \ "\i\Basis. (\a. a *\<^sub>R i) ` {x\i - d, x\i + d}" show "finite c" unfolding c_def by (simp add: finite_set_setsum) - have 1: "convex hull c = {a. \i\Basis. a \ i \ {x \ i - d..x \ i + d}}" + have 1: "convex hull c = {a. \i\Basis. a \ i \ cbox (x \ i - d) (x \ i + d)}" unfolding box_eq_set_setsum_Basis unfolding c_def convex_hull_set_setsum apply (subst convex_hull_linear_image [symmetric]) apply (simp add: linear_iff scaleR_add_left) apply (rule setsum_cong [OF refl]) apply (rule image_cong [OF _ refl]) - apply (rule convex_hull_eq_real_interval) + apply (rule convex_hull_eq_real_cbox) apply (cut_tac `0 < d`, simp) done then have 2: "convex hull c = {a. \i\Basis. a \ i \ cball (x \ i) d}" diff -r 2666cd7d380c -r 0268784f60da src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Mon Mar 17 21:56:32 2014 +0100 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Tue Mar 18 10:12:57 2014 +0100 @@ -458,11 +458,11 @@ by (rule has_derivative_unique) lemma frechet_derivative_unique_within_closed_interval: - fixes f::"'a::ordered_euclidean_space \ 'b::real_normed_vector" + fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes "\i\Basis. a\i < b\i" - and "x \ {a..b}" - and "(f has_derivative f' ) (at x within {a..b})" - and "(f has_derivative f'') (at x within {a..b})" + and "x \ cbox a b" + and "(f has_derivative f' ) (at x within cbox a b)" + and "(f has_derivative f'') (at x within cbox a b)" shows "f' = f''" apply(rule frechet_derivative_unique_within) apply(rule assms(3,4))+ @@ -470,18 +470,18 @@ fix e :: real fix i :: 'a assume "e > 0" and i: "i \ Basis" - then show "\d. 0 < \d\ \ \d\ < e \ x + d *\<^sub>R i \ {a..b}" + then show "\d. 0 < \d\ \ \d\ < e \ x + d *\<^sub>R i \ cbox a b" proof (cases "x\i = a\i") case True then show ?thesis apply (rule_tac x="(min (b\i - a\i) e) / 2" in exI) using assms(1)[THEN bspec[where x=i]] and `e>0` and assms(2) - unfolding mem_interval + unfolding mem_box using i apply (auto simp add: field_simps inner_simps inner_Basis) done next - note * = assms(2)[unfolded mem_interval, THEN bspec, OF i] + note * = assms(2)[unfolded mem_box, THEN bspec, OF i] case False moreover have "a \ i < x \ i" using False * by auto @@ -502,7 +502,7 @@ ultimately show ?thesis apply (rule_tac x="- (min (x\i - a\i) e) / 2" in exI) using assms(1)[THEN bspec, OF i] and `e>0` and assms(2) - unfolding mem_interval + unfolding mem_box using i apply (auto simp add: field_simps inner_simps inner_Basis) done @@ -510,14 +510,14 @@ qed lemma frechet_derivative_unique_within_open_interval: - fixes f::"'a::ordered_euclidean_space \ 'b::real_normed_vector" + fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes "x \ box a b" and "(f has_derivative f' ) (at x within box a b)" and "(f has_derivative f'') (at x within box a b)" shows "f' = f''" proof - from assms(1) have *: "at x within box a b = at x" - by (metis at_within_interior interior_open open_interval) + by (metis at_within_interior interior_open open_box) from assms(2,3) [unfolded *] show "f' = f''" by (rule frechet_derivative_unique_at) qed @@ -531,12 +531,12 @@ apply auto done -lemma frechet_derivative_within_closed_interval: - fixes f :: "'a::ordered_euclidean_space \ 'b::real_normed_vector" +lemma frechet_derivative_within_cbox: + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "\i\Basis. a\i < b\i" - and "x \ {a..b}" - and "(f has_derivative f') (at x within {a..b})" - shows "frechet_derivative f (at x within {a..b}) = f'" + and "x \ cbox a b" + and "(f has_derivative f') (at x within cbox a b)" + shows "frechet_derivative f (at x within cbox a b) = f'" using assms by (metis Derivative.differentiableI frechet_derivative_unique_within_closed_interval frechet_derivative_works) @@ -641,34 +641,34 @@ fixes f :: "real \ real" assumes "a < b" and "f a = f b" - and "continuous_on {a..b} f" - and "\x\box a b. (f has_derivative f' x) (at x)" - shows "\x\box a b. f' x = (\v. 0)" + and "continuous_on {a .. b} f" + and "\x\{a <..< b}. (f has_derivative f' x) (at x)" + shows "\x\{a <..< b}. f' x = (\v. 0)" proof - have "\x\box a b. (\y\box a b. f x \ f y) \ (\y\box a b. f y \ f x)" proof - have "(a + b) / 2 \ {a .. b}" using assms(1) by auto - then have *: "{a..b} \ {}" + then have *: "{a .. b} \ {}" by auto obtain d where d: - "d \ {a..b}" - "\y\{a..b}. f y \ f d" - using continuous_attains_sup[OF compact_interval * assms(3)] .. + "d \cbox a b" + "\y\cbox a b. f y \ f d" + using continuous_attains_sup[OF compact_Icc * assms(3)] by auto obtain c where c: - "c \ {a..b}" - "\y\{a..b}. f c \ f y" - using continuous_attains_inf[OF compact_interval * assms(3)] .. + "c \ cbox a b" + "\y\cbox a b. f c \ f y" + using continuous_attains_inf[OF compact_Icc * assms(3)] by auto show ?thesis proof (cases "d \ box a b \ c \ box a b") case True then show ?thesis - by (metis c(2) d(2) interval_open_subset_closed subset_iff) + by (metis c(2) d(2) box_subset_cbox subset_iff) next def e \ "(a + b) /2" case False then have "f d = f c" - using d c assms(2) by (auto simp: box_real) + using d c assms(2) by auto then have "\x. x \ {a..b} \ f x = f d" using c d by force @@ -676,11 +676,12 @@ apply (rule_tac x=e in bexI) unfolding e_def using assms(1) - apply (auto simp: box_real) + apply auto done qed qed - then obtain x where x: "x \ box a b" "(\y\box a b. f x \ f y) \ (\y\box a b. f y \ f x)" .. + then obtain x where x: "x \ {a <..< b}" "(\y\{a <..< b}. f x \ f y) \ (\y\{a <..< b}. f y \ f x)" + by auto then have "f' x = (\v. 0)" apply (rule_tac differential_zero_maxmin[of x "box a b" f "f' x"]) using assms @@ -700,19 +701,21 @@ assumes "\x\{a<..x\{a<..x\box a b. (\xa. f' x xa - (f b - f a) / (b - a) * xa) = (\v. 0)" + have "\x\{a <..< b}. (\xa. f' x xa - (f b - f a) / (b - a) * xa) = (\v. 0)" proof (intro rolle[OF assms(1), of "\x. f x - (f b - f a) / (b - a) * x"] ballI) fix x - assume "x \ box a b" hence x: "x \ {a<.. {a <..< b}" show "((\x. f x - (f b - f a) / (b - a) * x) has_derivative (\xa. f' x xa - (f b - f a) / (b - a) * xa)) (at x)" by (intro has_derivative_intros assms(3)[rule_format,OF x] mult_right_has_derivative) qed (insert assms(1,2), auto intro!: continuous_on_intros simp: field_simps) then obtain x where - "x \ box a b" + "x \ {a <..< b}" "(\xa. f' x xa - (f b - f a) / (b - a) * xa) = (\v. 0)" .. then show ?thesis - by (metis (erased, hide_lams) assms(1) box_real diff_less_iff(1) eq_iff_diff_eq_0 linordered_field_class.sign_simps(41) nonzero_mult_divide_cancel_right not_real_square_gt_zero times_divide_eq_left) + by (metis (erased, hide_lams) assms(1) diff_less_iff(1) eq_iff_diff_eq_0 + linordered_field_class.sign_simps(41) nonzero_mult_divide_cancel_right not_real_square_gt_zero + times_divide_eq_left) qed lemma mvt_simple: @@ -727,21 +730,21 @@ defer proof fix x - assume x: "x \ {a<.. box a b" by (simp add: box_real) + assume x: "x \ {a <..< b}" show "(f has_derivative f' x) (at x)" - unfolding has_derivative_within_open[OF x open_interval,symmetric] + unfolding has_derivative_within_open[OF x open_greaterThanLessThan,symmetric] apply (rule has_derivative_within_subset) apply (rule assms(2)[rule_format]) using x - apply (auto simp: box_real) + apply auto done qed (insert assms(2), auto) lemma mvt_very_simple: fixes f :: "real \ real" assumes "a \ b" - and "\x\{a..b}. (f has_derivative f' x) (at x within {a..b})" - shows "\x\{a..b}. f b - f a = f' x (b - a)" + and "\x\{a .. b}. (f has_derivative f' x) (at x within {a .. b})" + shows "\x\{a .. b}. f b - f a = f' x (b - a)" proof (cases "a = b") interpret bounded_linear "f' b" using assms(2) assms(1) by auto @@ -767,7 +770,7 @@ lemma mvt_general: fixes f :: "real \ 'a::euclidean_space" assumes "a < b" - and "continuous_on {a..b} f" + and "continuous_on {a .. b} f" and "\x\{a<..x\{a<.. norm (f' x (b - a))" proof - @@ -825,7 +828,7 @@ using assms(1)[unfolded convex_alt,rule_format,OF x y] unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib by (auto simp add: algebra_simps) - then have 1: "continuous_on {0..1} (f \ ?p)" + then have 1: "continuous_on {0 .. 1} (f \ ?p)" apply - apply (rule continuous_on_intros)+ unfolding continuous_on_eq_continuous_within @@ -837,12 +840,12 @@ apply (rule assms(2)[rule_format]) apply auto done - have 2: "\u\{0<..<1}. + have 2: "\u\{0 <..< 1}. ((f \ ?p) has_derivative f' (x + u *\<^sub>R (y - x)) \ (\u. 0 + u *\<^sub>R (y - x))) (at u)" proof rule case goal1 let ?u = "x + u *\<^sub>R (y - x)" - have "(f \ ?p has_derivative (f' ?u) \ (\u. 0 + u *\<^sub>R (y - x))) (at u within {0<..<1})" + have "(f \ ?p has_derivative (f' ?u) \ (\u. 0 + u *\<^sub>R (y - x))) (at u within box 0 1)" apply (rule diff_chain_within) apply (rule has_derivative_intros)+ apply (rule has_derivative_within_subset) @@ -851,7 +854,7 @@ apply auto done then show ?case - unfolding has_derivative_within_open[OF goal1 open_greaterThanLessThan] . + by (simp add: has_derivative_within_open[OF goal1 open_greaterThanLessThan]) qed obtain u where u: "u \ {0<..<1}" @@ -2051,9 +2054,9 @@ lemma vector_derivative_unique_within_closed_interval: assumes "a < b" - and "x \ {a..b}" - assumes "(f has_vector_derivative f') (at x within {a..b})" - assumes "(f has_vector_derivative f'') (at x within {a..b})" + and "x \ cbox a b" + assumes "(f has_vector_derivative f') (at x within cbox a b)" + assumes "(f has_vector_derivative f'') (at x within cbox a b)" shows "f' = f''" proof - have *: "(\x. x *\<^sub>R f') = (\x. x *\<^sub>R f'')" @@ -2084,9 +2087,9 @@ lemma vector_derivative_within_closed_interval: assumes "a < b" - and "x \ {a..b}" - assumes "(f has_vector_derivative f') (at x within {a..b})" - shows "vector_derivative f (at x within {a..b}) = f'" + and "x \ cbox a b" + assumes "(f has_vector_derivative f') (at x within cbox a b)" + shows "vector_derivative f (at x within cbox a b) = f'" apply (rule vector_derivative_unique_within_closed_interval) using vector_derivative_works[unfolded differentiable_def] using assms diff -r 2666cd7d380c -r 0268784f60da src/HOL/Multivariate_Analysis/Fashoda.thy --- a/src/HOL/Multivariate_Analysis/Fashoda.thy Mon Mar 17 21:56:32 2014 +0100 +++ b/src/HOL/Multivariate_Analysis/Fashoda.thy Tue Mar 18 10:12:57 2014 +0100 @@ -38,14 +38,14 @@ lemma fashoda_unit: fixes f g :: "real \ real^2" - assumes "f ` {- 1..1} \ {- 1..1}" - and "g ` {- 1..1} \ {- 1..1}" - and "continuous_on {- 1..1} f" - and "continuous_on {- 1..1} g" + assumes "f ` {-1 .. 1} \ cbox (-1) 1" + and "g ` {-1 .. 1} \ cbox (-1) 1" + and "continuous_on {-1 .. 1} f" + and "continuous_on {-1 .. 1} g" and "f (- 1)$1 = - 1" and "f 1$1 = 1" "g (- 1) $2 = -1" and "g 1 $2 = 1" - shows "\s\{- 1..1}. \t\{- 1..1}. f s = g t" + shows "\s\{-1 .. 1}. \t\{-1 .. 1}. f s = g t" proof (rule ccontr) assume "\ ?thesis" note as = this[unfolded bex_simps,rule_format] @@ -61,9 +61,9 @@ apply auto done let ?F = "\w::real^2. (f \ (\x. x$1)) w - (g \ (\x. x$2)) w" - have *: "\i. (\x::real^2. x $ i) ` {- 1..1} = {- 1..1::real}" + have *: "\i. (\x::real^2. x $ i) ` cbox (- 1) 1 = {-1 .. 1}" apply (rule set_eqI) - unfolding image_iff Bex_def mem_interval_cart + unfolding image_iff Bex_def mem_interval_cart interval_cbox_cart apply rule defer apply (rule_tac x="vec x" in exI) @@ -71,21 +71,21 @@ done { fix x - assume "x \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w) ` {- 1..1::real^2}" + assume "x \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w) ` (cbox (- 1) (1::real^2))" then obtain w :: "real^2" where w: - "w \ {- 1..1}" + "w \ cbox (- 1) 1" "x = (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w" unfolding image_iff .. then have "x \ 0" using as[of "w$1" "w$2"] - unfolding mem_interval_cart + unfolding mem_interval_cart atLeastAtMost_iff by auto } note x0 = this have 21: "\i::2. i \ 1 \ i = 2" using UNIV_2 by auto have 1: "box (- 1) (1::real^2) \ {}" unfolding interval_eq_empty_cart by auto - have 2: "continuous_on {- 1..1} (negatex \ sqprojection \ ?F)" + have 2: "continuous_on (cbox -1 1) (negatex \ sqprojection \ ?F)" apply (intro continuous_on_intros continuous_on_component) unfolding * apply (rule assms)+ @@ -112,13 +112,13 @@ done qed qed - have 3: "(negatex \ sqprojection \ ?F) ` {- 1..1} \ {- 1..1}" + have 3: "(negatex \ sqprojection \ ?F) ` cbox (-1) 1 \ cbox (-1) 1" unfolding subset_eq apply rule proof - case goal1 then obtain y :: "real^2" where y: - "y \ {- 1..1}" + "y \ cbox -1 1" "x = (negatex \ sqprojection \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w)) y" unfolding image_iff .. have "?F y \ 0" @@ -132,8 +132,8 @@ have "infnorm x = 1" unfolding *[symmetric] y o_def by (rule lem1[rule_format]) - then show "x \ {- 1..1}" - unfolding mem_interval_cart infnorm_2 + then show "x \ cbox (-1) 1" + unfolding mem_interval_cart interval_cbox_cart infnorm_2 apply - apply rule proof - @@ -147,11 +147,11 @@ qed qed obtain x :: "real^2" where x: - "x \ {- 1..1}" + "x \ cbox -1 1" "(negatex \ sqprojection \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w)) x = x" - apply (rule brouwer_weak[of "{- 1..1::real^2}" "negatex \ sqprojection \ ?F"]) - apply (rule compact_interval convex_interval)+ - unfolding interior_closed_interval + apply (rule brouwer_weak[of "cbox -1 (1::real^2)" "negatex \ sqprojection \ ?F"]) + apply (rule compact_cbox convex_box)+ + unfolding interior_cbox apply (rule 1 2 3)+ apply blast done @@ -213,7 +213,7 @@ unfolding as negatex_def vector_2 by auto moreover - from x1 have "g (x $ 2) \ {- 1..1}" + from x1 have "g (x $ 2) \ cbox (-1) 1" apply - apply (rule assms(2)[unfolded subset_eq,rule_format]) apply auto @@ -232,7 +232,7 @@ unfolding as negatex_def vector_2 by auto moreover - from x1 have "g (x $ 2) \ {- 1..1}" + from x1 have "g (x $ 2) \ cbox (-1) 1" apply - apply (rule assms(2)[unfolded subset_eq,rule_format]) apply auto @@ -251,7 +251,7 @@ unfolding as negatex_def vector_2 by auto moreover - from x1 have "f (x $ 1) \ {- 1..1}" + from x1 have "f (x $ 1) \ cbox (-1) 1" apply - apply (rule assms(1)[unfolded subset_eq,rule_format]) apply auto @@ -270,7 +270,7 @@ unfolding as negatex_def vector_2 by auto moreover - from x1 have "f (x $ 1) \ {- 1..1}" + from x1 have "f (x $ 1) \ cbox (-1) 1" apply - apply (rule assms(1)[unfolded subset_eq,rule_format]) apply auto @@ -287,8 +287,8 @@ fixes f g :: "real \ real^2" assumes "path f" and "path g" - and "path_image f \ {- 1..1}" - and "path_image g \ {- 1..1}" + and "path_image f \ cbox (-1) 1" + and "path_image g \ cbox (-1) 1" and "(pathstart f)$1 = -1" and "(pathfinish f)$1 = 1" and "(pathstart g)$2 = -1" @@ -301,7 +301,7 @@ unfolding iscale_def by auto have "\s\{- 1..1}. \t\{- 1..1}. (f \ iscale) s = (g \ iscale) t" proof (rule fashoda_unit) - show "(f \ iscale) ` {- 1..1} \ {- 1..1}" "(g \ iscale) ` {- 1..1} \ {- 1..1}" + show "(f \ iscale) ` {- 1..1} \ cbox -1 1" "(g \ iscale) ` {- 1..1} \ cbox -1 1" using isc and assms(3-4) by (auto simp add: image_comp [symmetric]) have *: "continuous_on {- 1..1} iscale" unfolding iscale_def by (rule continuous_on_intros)+ @@ -325,7 +325,7 @@ "s \ {- 1..1}" "t \ {- 1..1}" "(f \ iscale) s = (g \ iscale) t" - by blast + by auto show thesis apply (rule_tac z = "f (iscale s)" in that) using st @@ -343,8 +343,8 @@ fixes b :: "real^2" assumes "path f" and "path g" - and "path_image f \ {a..b}" - and "path_image g \ {a..b}" + and "path_image f \ cbox a b" + and "path_image g \ cbox a b" and "(pathstart f)$1 = a$1" and "(pathfinish f)$1 = b$1" and "(pathstart g)$2 = a$2" @@ -356,7 +356,7 @@ then show thesis by auto next - have "{a..b} \ {}" + have "cbox a b \ {}" using assms(3) using path_image_nonempty[of f] by auto then have "a \ b" unfolding interval_eq_empty_cart less_eq_vec_def by (auto simp add: not_less) @@ -371,13 +371,13 @@ apply (rule pathfinish_in_path_image) unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"] unfolding pathstart_def - apply (auto simp add: less_eq_vec_def) + apply (auto simp add: less_eq_vec_def mem_interval_cart) done then obtain z :: "real^2" where z: "z \ path_image g" "z $ 2 = pathstart f $ 2" .. - have "z \ {a..b}" + have "z \ cbox a b" using z(1) assms(4) unfolding path_image_def - by blast + by blast then have "z = f 0" unfolding vec_eq_iff forall_2 unfolding z(2) pathstart_def @@ -403,13 +403,13 @@ unfolding assms using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"] unfolding pathstart_def - apply (auto simp add: less_eq_vec_def) + apply (auto simp add: less_eq_vec_def mem_interval_cart) done then obtain z where z: "z \ path_image f" "z $ 1 = pathstart g $ 1" .. - have "z \ {a..b}" + have "z \ cbox a b" using z(1) assms(3) unfolding path_image_def - by blast + by blast then have "z = g 0" unfolding vec_eq_iff forall_2 unfolding z(2) pathstart_def @@ -427,7 +427,7 @@ done next assume as: "a $ 1 < b $ 1 \ a $ 2 < b $ 2" - have int_nem: "{- 1..1::real^2} \ {}" + have int_nem: "cbox (-1) (1::real^2) \ {}" unfolding interval_eq_empty_cart by auto obtain z :: "real^2" where z: "z \ (interval_bij (a, b) (- 1, 1) \ f) ` {0..1}" @@ -445,7 +445,7 @@ "y \ {0..1}" "x = (interval_bij (a, b) (- 1, 1) \ f) y" unfolding image_iff .. - show "x \ {- 1..1}" + show "x \ cbox -1 1" unfolding y o_def apply (rule in_interval_interval_bij) using y(1) @@ -459,7 +459,7 @@ "y \ {0..1}" "x = (interval_bij (a, b) (- 1, 1) \ g) y" unfolding image_iff .. - show "x \ {- 1..1}" + show "x \ cbox -1 1" unfolding y o_def apply (rule in_interval_interval_bij) using y(1) @@ -471,7 +471,7 @@ and "(interval_bij (a, b) (- 1, 1) \ f) 1 $ 1 = 1" and "(interval_bij (a, b) (- 1, 1) \ g) 0 $ 2 = -1" and "(interval_bij (a, b) (- 1, 1) \ g) 1 $ 2 = 1" - using assms as + using assms as by (simp_all add: axis_in_Basis cart_eq_inner_axis pathstart_def pathfinish_def interval_bij_def) (simp_all add: inner_axis) qed @@ -669,8 +669,8 @@ fixes a :: "real^2" assumes "path f" and "path g" - and "path_image f \ {a..b}" - and "path_image g \ {a..b}" + and "path_image f \ cbox a b" + and "path_image g \ cbox a b" and "(pathstart f)$2 = a$2" and "(pathfinish f)$2 = a$2" and "(pathstart g)$2 = a$2" @@ -680,13 +680,13 @@ and "(pathfinish f)$1 < (pathfinish g)$1" obtains z where "z \ path_image f" and "z \ path_image g" proof - - have "{a..b} \ {}" + have "cbox a b \ {}" using path_image_nonempty[of f] using assms(3) by auto note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less] - have "pathstart f \ {a..b}" - and "pathfinish f \ {a..b}" - and "pathstart g \ {a..b}" - and "pathfinish g \ {a..b}" + have "pathstart f \ cbox a b" + and "pathfinish f \ cbox a b" + and "pathstart g \ cbox a b" + and "pathfinish g \ cbox a b" using pathstart_in_path_image pathfinish_in_path_image using assms(3-4) by auto @@ -710,7 +710,8 @@ path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \ path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2) by(auto simp add: path_image_join path_linepath) - have abab: "{a..b} \ {?a..?b}" + have abab: "cbox a b \ cbox ?a ?b" + unfolding interval_cbox_cart[symmetric] by (auto simp add:less_eq_vec_def forall_2 vector_2) obtain z where "z \ path_image @@ -730,30 +731,30 @@ proof - show "path ?P1" and "path ?P2" using assms by auto - have "path_image ?P1 \ {?a .. ?b}" + have "path_image ?P1 \ cbox ?a ?b" unfolding P1P2 path_image_linepath apply (rule Un_least)+ defer 3 - apply (rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format]) + apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) unfolding mem_interval_cart forall_2 vector_2 using ab startfin abab assms(3) using assms(9-) unfolding assms - apply (auto simp add: field_simps interval) + apply (auto simp add: field_simps box) done - then show "path_image ?P1 \ {?a .. ?b}" . - have "path_image ?P2 \ {?a .. ?b}" + then show "path_image ?P1 \ cbox ?a ?b" . + have "path_image ?P2 \ cbox ?a ?b" unfolding P1P2 path_image_linepath apply (rule Un_least)+ defer 2 - apply (rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format]) + apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) unfolding mem_interval_cart forall_2 vector_2 using ab startfin abab assms(4) using assms(9-) unfolding assms - apply (auto simp add: field_simps interval) + apply (auto simp add: field_simps box) done - then show "path_image ?P2 \ {?a .. ?b}" . + then show "path_image ?P2 \ cbox ?a ?b" . show "a $ 1 - 2 = a $ 1 - 2" and "b $ 1 + 2 = b $ 1 + 2" and "pathstart g $ 2 - 3 = a $ 2 - 3" @@ -775,7 +776,7 @@ apply (simp only: segment_vertical segment_horizontal vector_2) proof - case goal1 note as=this - have "pathfinish f \ {a..b}" + have "pathfinish f \ cbox a b" using assms(3) pathfinish_in_path_image[of f] by auto then have "1 + b $ 1 \ pathfinish f $ 1 \ False" unfolding mem_interval_cart forall_2 by auto @@ -783,7 +784,7 @@ using as(2) using assms ab by (auto simp add: field_simps) - moreover have "pathstart f \ {a..b}" + moreover have "pathstart f \ cbox a b" using assms(3) pathstart_in_path_image[of f] by auto then have "1 + b $ 1 \ pathstart f $ 1 \ False" @@ -799,7 +800,7 @@ using as(2) using assms ab by (auto simp add: field_simps *) - moreover have "pathstart g \ {a..b}" + moreover have "pathstart g \ cbox a b" using assms(4) pathstart_in_path_image[of g] by auto note this[unfolded mem_interval_cart forall_2] @@ -816,7 +817,7 @@ qed then have "z \ path_image f \ z \ path_image g" using z unfolding Un_iff by blast - then have z': "z \ {a..b}" + then have z': "z \ cbox a b" using assms(3-4) by auto have "a $ 2 = z $ 2 \ (z $ 1 = pathstart f $ 1 \ z $ 1 = pathfinish f $ 1) \ diff -r 2666cd7d380c -r 0268784f60da src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Mon Mar 17 21:56:32 2014 +0100 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Mar 18 10:12:57 2014 +0100 @@ -219,26 +219,27 @@ abbreviation One :: "'a::euclidean_space" where "One \ \Basis" -lemma empty_as_interval: "{} = {One..(0::'a::ordered_euclidean_space)}" - by (auto simp: eucl_le[where 'a='a]) +lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)" + using nonempty_Basis + by (fastforce simp add: set_eq_iff mem_box) lemma One_nonneg: "0 \ (One::'a::ordered_euclidean_space)" by (auto intro: setsum_nonneg) lemma interior_subset_union_intervals: - assumes "i = {a..b::'a::ordered_euclidean_space}" - and "j = {c..d}" + assumes "i = cbox a b" + and "j = cbox c d" and "interior j \ {}" and "i \ j \ s" and "interior i \ interior j = {}" shows "interior i \ interior s" proof - - have "box a b \ {c..d} = {}" - using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5) - unfolding assms(1,2) interior_closed_interval by auto + have "box a b \ cbox c d = {}" + using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5) + unfolding assms(1,2) interior_cbox by auto moreover - have "box a b \ {c..d} \ s" - apply (rule order_trans,rule interval_open_subset_closed) + have "box a b \ cbox c d \ s" + apply (rule order_trans,rule box_subset_cbox) using assms(4) unfolding assms(1,2) apply auto done @@ -248,16 +249,16 @@ apply (rule interior_maximal) defer apply (rule open_interior) - unfolding assms(1,2) interior_closed_interval + unfolding assms(1,2) interior_cbox apply auto done qed lemma inter_interior_unions_intervals: - fixes f::"('a::ordered_euclidean_space) set set" + fixes f::"('a::euclidean_space) set set" assumes "finite f" and "open s" - and "\t\f. \a b. t = {a..b}" + and "\t\f. \a b. t = cbox a b" and "\t\f. s \ (interior t) = {}" shows "s \ interior (\f) = {}" proof (rule ccontr, unfold ex_in_conv[symmetric]) @@ -271,7 +272,7 @@ apply auto done have lem2: "\x s P. \x\s. P x \ \x\insert x s. P x" by auto - have "\f. finite f \ \t\f. \a b. t = {a..b} \ + have "\f. finite f \ \t\f. \a b. t = cbox a b \ \x. x \ s \ interior (\f) \ \t\f. \x. \e>0. ball x e \ s \ t" proof - case goal1 @@ -289,16 +290,16 @@ using insert(5) .. then obtain e where e: "0 < e \ ball x e \ s \ interior (\insert i f)" unfolding open_contains_ball_eq[OF open_Int[OF assms(2) open_interior], rule_format] .. - obtain a where "\b. i = {a..b}" + obtain a where "\b. i = cbox a b" using insert(4)[rule_format,OF insertI1] .. - then obtain b where ab: "i = {a..b}" .. + then obtain b where ab: "i = cbox a b" .. show ?case proof (cases "x \ i") case False - then have "x \ UNIV - {a..b}" + then have "x \ UNIV - cbox a b" unfolding ab by auto - then obtain d where "0 < d \ ball x d \ UNIV - {a..b}" - unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_interval],rule_format] .. + then obtain d where "0 < d \ ball x d \ UNIV - cbox a b" + unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_cbox],rule_format] .. then have "0 < d" "ball x (min d e) \ UNIV - i" unfolding ab ball_min_Int by auto then have "ball x (min d e) \ s \ interior (\f)" @@ -316,19 +317,19 @@ proof (cases "x\box a b") case True then obtain d where "0 < d \ ball x d \ box a b" - unfolding open_contains_ball_eq[OF open_interval,rule_format] .. + unfolding open_contains_ball_eq[OF open_box,rule_format] .. then show ?thesis apply (rule_tac x=i in bexI, rule_tac x=x in exI, rule_tac x="min d e" in exI) unfolding ab - using interval_open_subset_closed[of a b] and e + using box_subset_cbox[of a b] and e apply fastforce+ done next case False then obtain k where "x\k \ a\k \ x\k \ b\k" and k: "k \ Basis" - unfolding mem_interval by (auto simp add: not_less) + unfolding mem_box by (auto simp add: not_less) then have "x\k = a\k \ x\k = b\k" - using True unfolding ab and mem_interval + using True unfolding ab and mem_box apply (erule_tac x = k in ballE) apply auto done @@ -350,7 +351,7 @@ using e[THEN conjunct1] k by (auto simp add: field_simps abs_less_iff as inner_Basis inner_simps) then have "y \ i" - unfolding ab mem_interval by (auto intro!: bexI[OF _ k]) + unfolding ab mem_box by (auto intro!: bexI[OF _ k]) then show False using yi by auto qed moreover @@ -399,7 +400,7 @@ using e[THEN conjunct1] k by (auto simp add:field_simps inner_simps inner_Basis as) then have "y \ i" - unfolding ab mem_interval by (auto intro!: bexI[OF _ k]) + unfolding ab mem_box by (auto intro!: bexI[OF _ k]) then show False using yi by auto qed moreover @@ -452,94 +453,141 @@ using `t \ f` assms(4) by auto qed -lemma interval_bounds: - fixes a b::"'a::ordered_euclidean_space" - shows "\i\Basis. a\i \ b\i \ Inf {a..b} = a" "\i\Basis. a\i \ b\i \ Sup {a..b} = b" - by (auto simp: eucl_le[where 'a='a]) - -lemma interval_bounds': - fixes a b::"'a::ordered_euclidean_space" - assumes "{a..b} \ {}" - shows "Sup {a..b} = b" - and "Inf {a..b} = a" - using assms - by (auto simp: eucl_le[where 'a='a]) +subsection {* Bounds on intervals where they exist. *} + +definition interval_upperbound :: "('a::euclidean_space) set \ 'a" + where "interval_upperbound s = (\i\Basis. (SUP x:s. x\i) *\<^sub>R i)" + +definition interval_lowerbound :: "('a::euclidean_space) set \ 'a" + where "interval_lowerbound s = (\i\Basis. (INF x:s. x\i) *\<^sub>R i)" + +lemma interval_upperbound[simp]: + "\i\Basis. a\i \ b\i \ + interval_upperbound (cbox a b) = (b::'a::euclidean_space)" + unfolding interval_upperbound_def euclidean_representation_setsum cbox_def SUP_def + by (safe intro!: cSup_eq) auto + +lemma interval_lowerbound[simp]: + "\i\Basis. a\i \ b\i \ + interval_lowerbound (cbox a b) = (a::'a::euclidean_space)" + unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def INF_def + by (safe intro!: cInf_eq) auto + +lemmas interval_bounds = interval_upperbound interval_lowerbound + +lemma + fixes X::"real set" + shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X" + and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X" + by (auto simp: interval_upperbound_def interval_lowerbound_def SUP_def INF_def) + +lemma interval_bounds'[simp]: + assumes "cbox a b \ {}" + shows "interval_upperbound (cbox a b) = b" + and "interval_lowerbound (cbox a b) = a" + using assms unfolding box_ne_empty by auto subsection {* Content (length, area, volume...) of an interval. *} -definition "content (s::('a::ordered_euclidean_space) set) = - (if s = {} then 0 else (\i\Basis. (Sup s)\i - (Inf s)\i))" - -lemma interval_not_empty: "\i\Basis. a\i \ b\i \ {a..b::'a::ordered_euclidean_space} \ {}" - unfolding interval_eq_empty unfolding not_ex not_less by auto - -lemma content_closed_interval: - fixes a :: "'a::ordered_euclidean_space" +definition "content (s::('a::euclidean_space) set) = + (if s = {} then 0 else (\i\Basis. (interval_upperbound s)\i - (interval_lowerbound s)\i))" + +lemma interval_not_empty: "\i\Basis. a\i \ b\i \ cbox a b \ {}" + unfolding box_eq_empty unfolding not_ex not_less by auto + +lemma content_cbox: + fixes a :: "'a::euclidean_space" assumes "\i\Basis. a\i \ b\i" - shows "content {a..b} = (\i\Basis. b\i - a\i)" + shows "content (cbox a b) = (\i\Basis. b\i - a\i)" using interval_not_empty[OF assms] unfolding content_def - by auto - -lemma content_closed_interval': - fixes a :: "'a::ordered_euclidean_space" - assumes "{a..b} \ {}" - shows "content {a..b} = (\i\Basis. b\i - a\i)" - apply (rule content_closed_interval) + by (auto simp: ) + +lemma content_cbox': + fixes a :: "'a::euclidean_space" + assumes "cbox a b \ {}" + shows "content (cbox a b) = (\i\Basis. b\i - a\i)" + apply (rule content_cbox) using assms - unfolding interval_ne_empty + unfolding box_ne_empty apply assumption done lemma content_real: "a \ b \ content {a..b} = b - a" - unfolding content_def by auto + by (auto simp: interval_upperbound_def interval_lowerbound_def SUP_def INF_def content_def) + +lemma content_closed_interval: + fixes a :: "'a::ordered_euclidean_space" + assumes "a \ b" + shows "content {a .. b} = (\i\Basis. b\i - a\i)" + using content_cbox[of a b] assms + by (simp add: cbox_interval eucl_le[where 'a='a]) lemma content_singleton[simp]: "content {a} = 0" proof - - have "content {a .. a} = 0" - by (subst content_closed_interval) (auto simp: ex_in_conv) - then show ?thesis by simp -qed - -lemma content_unit[intro]: "content{0..One::'a::ordered_euclidean_space} = 1" - by (auto simp: content_def eucl_le[where 'a='a]) + have "content (cbox a a) = 0" + by (subst content_cbox) (auto simp: ex_in_conv) + then show ?thesis by (simp add: cbox_sing) +qed + +lemma content_unit[intro]: "content(cbox 0 (One::'a::euclidean_space)) = 1" + proof - + have *: "\i\Basis. (0::'a)\i \ (One::'a)\i" + by auto + have "0 \ cbox 0 (One::'a)" + unfolding mem_box by auto + then show ?thesis + unfolding content_def interval_bounds[OF *] using setprod_1 by auto + qed lemma content_pos_le[intro]: - fixes a::"'a::ordered_euclidean_space" - shows "0 \ content {a..b}" - by (auto simp: content_def eucl_le[where 'a='a] intro!: setprod_nonneg) + fixes a::"'a::euclidean_space" + shows "0 \ content (cbox a b)" +proof (cases "cbox a b = {}") + case False + then have *: "\i\Basis. a \ i \ b \ i" + unfolding box_ne_empty . + have "0 \ (\i\Basis. interval_upperbound (cbox a b) \ i - interval_lowerbound (cbox a b) \ i)" + apply (rule setprod_nonneg) + unfolding interval_bounds[OF *] + using * + apply auto + done + also have "\ = content (cbox a b)" using False by (simp add: content_def) + finally show ?thesis . +qed (simp add: content_def) lemma content_pos_lt: - fixes a :: "'a::ordered_euclidean_space" + fixes a :: "'a::euclidean_space" assumes "\i\Basis. a\i < b\i" - shows "0 < content {a..b}" + shows "0 < content (cbox a b)" using assms - by (auto simp: content_def eucl_le[where 'a='a] intro!: setprod_pos) + by (auto simp: content_def box_eq_empty intro!: setprod_pos) lemma content_eq_0: - "content{a..b::'a::ordered_euclidean_space} = 0 \ (\i\Basis. b\i \ a\i)" - by (auto intro!: bexI simp: content_def eucl_le[where 'a='a]) + "content (cbox a b) = 0 \ (\i\Basis. b\i \ a\i)" + by (auto simp: content_def box_eq_empty intro!: setprod_pos bexI) lemma cond_cases: "(P \ Q x) \ (\ P \ Q y) \ Q (if P then x else y)" by auto -lemma content_closed_interval_cases: - "content {a..b::'a::ordered_euclidean_space} = +lemma content_cbox_cases: + "content (cbox a (b::'a::euclidean_space)) = (if \i\Basis. a\i \ b\i then setprod (\i. b\i - a\i) Basis else 0)" - by (auto simp: not_le content_eq_0 intro: less_imp_le content_closed_interval) - -lemma content_eq_0_interior: "content {a..b} = 0 \ interior({a..b}) = {}" - unfolding content_eq_0 interior_closed_interval interval_eq_empty + by (auto simp: not_le content_eq_0 intro: less_imp_le content_cbox) + +lemma content_eq_0_interior: "content (cbox a b) = 0 \ interior(cbox a b) = {}" + unfolding content_eq_0 interior_cbox box_eq_empty by auto lemma content_pos_lt_eq: - "0 < content {a..b::'a::ordered_euclidean_space} \ (\i\Basis. a\i < b\i)" + "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" apply rule defer apply (rule content_pos_lt, assumption) proof - - assume "0 < content {a..b}" - then have "content {a..b} \ 0" by auto + assume "0 < content (cbox a b)" + then have "content (cbox a b) \ 0" by auto then show "\i\Basis. a\i < b\i" unfolding content_eq_0 not_ex not_le by fastforce qed @@ -548,17 +596,40 @@ unfolding content_def by auto lemma content_subset: - assumes "{a..b} \ {c..d}" - shows "content {a..b::'a::ordered_euclidean_space} \ content {c..d}" -proof cases - assume "{a..b} \ {}" - with assms have "c \ a" "a \ b" "b \ d" "b + c \ d + a" by (auto intro: add_mono) - hence "b - a \ d - c" "c \ d" by (auto simp add: algebra_simps) - thus ?thesis - by (auto simp: content_def eucl_le[where 'a='a] inner_diff_left intro!: setprod_nonneg setprod_mono) -qed auto - -lemma content_lt_nz: "0 < content {a..b} \ content {a..b} \ 0" + assumes "cbox a b \ cbox c d" + shows "content (cbox a b) \ content (cbox c d)" +proof (cases "cbox a b = {}") + case True + then show ?thesis + using content_pos_le[of c d] by auto +next + case False + then have ab_ne: "\i\Basis. a \ i \ b \ i" + unfolding box_ne_empty by auto + then have ab_ab: "a\cbox a b" "b\cbox a b" + unfolding mem_box by auto + have "cbox c d \ {}" using assms False by auto + then have cd_ne: "\i\Basis. c \ i \ d \ i" + using assms unfolding box_ne_empty by auto + show ?thesis + unfolding content_def + unfolding interval_bounds[OF ab_ne] interval_bounds[OF cd_ne] + unfolding if_not_P[OF False] if_not_P[OF `cbox c d \ {}`] + apply (rule setprod_mono) + apply rule + proof + fix i :: 'a + assume i: "i \ Basis" + show "0 \ b \ i - a \ i" + using ab_ne[THEN bspec, OF i] i by auto + show "b \ i - a \ i \ d \ i - c \ i" + using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(2),of i] + using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(1),of i] + using i by auto + qed +qed + +lemma content_lt_nz: "0 < content (cbox a b) \ content (cbox a b) \ 0" unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce @@ -613,7 +684,7 @@ where "s division_of i \ finite s \ - (\k\s. k \ i \ k \ {} \ (\a b. k = {a..b})) \ + (\k\s. k \ i \ k \ {} \ (\a b. k = cbox a b)) \ (\k1\s. \k2\s. k1 \ k2 \ interior(k1) \ interior(k2) = {}) \ (\s = i)" @@ -622,7 +693,7 @@ shows "finite s" and "\k. k \ s \ k \ i" and "\k. k \ s \ k \ {}" - and "\k. k \ s \ \a b. k = {a..b}" + and "\k. k \ s \ \a b. k = cbox a b" and "\k1 k2. k1 \ s \ k2 \ s \ k1 \ k2 \ interior(k1) \ interior(k2) = {}" and "\s = i" using assms unfolding division_of_def by auto @@ -631,7 +702,7 @@ assumes "finite s" and "\k. k \ s \ k \ i" and "\k. k \ s \ k \ {}" - and "\k. k \ s \ \a b. k = {a..b}" + and "\k. k \ s \ \a b. k = cbox a b" and "\k1 k2. k1 \ s \ k2 \ s \ k1 \ k2 \ interior k1 \ interior k2 = {}" and "\s = i" shows "s division_of i" @@ -640,14 +711,14 @@ lemma division_of_finite: "s division_of i \ finite s" unfolding division_of_def by auto -lemma division_of_self[intro]: "{a..b} \ {} \ {{a..b}} division_of {a..b}" +lemma division_of_self[intro]: "cbox a b \ {} \ {cbox a b} division_of (cbox a b)" unfolding division_of_def by auto lemma division_of_trivial[simp]: "s division_of {} \ s = {}" unfolding division_of_def by auto lemma division_of_sing[simp]: - "s division_of {a..a::'a::ordered_euclidean_space} \ s = {{a..a}}" + "s division_of cbox a (a::'a::euclidean_space) \ s = {cbox a a}" (is "?l = ?r") proof assume ?r @@ -655,17 +726,17 @@ { assume "s = {{a}}" moreover fix k assume "k\s" - ultimately have"\x y. k = {x..y}" + ultimately have"\x y. k = cbox x y" apply (rule_tac x=a in exI)+ - unfolding interval_sing + unfolding cbox_sing apply auto done } ultimately show ?l - unfolding division_of_def interval_sing by auto + unfolding division_of_def cbox_sing by auto next assume ?l - note * = conjunctD4[OF this[unfolded division_of_def interval_sing]] + note * = conjunctD4[OF this[unfolded division_of_def cbox_sing]] { fix x assume x: "x \ s" have "x = {a}" @@ -674,20 +745,20 @@ moreover have "s \ {}" using *(4) by auto ultimately show ?r - unfolding interval_sing by auto + unfolding cbox_sing by auto qed lemma elementary_empty: obtains p where "p division_of {}" unfolding division_of_trivial by auto -lemma elementary_interval: obtains p where "p division_of {a..b}" +lemma elementary_interval: obtains p where "p division_of (cbox a b)" by (metis division_of_trivial division_of_self) lemma division_contains: "s division_of i \ \x\i. \k\s. x \ k" unfolding division_of_def by auto lemma forall_in_division: - "d division_of i \ (\x\d. P x) \ (\a b. {a..b} \ d \ P {a..b})" + "d division_of i \ (\x\d. P x) \ (\a b. cbox a b \ d \ P (cbox a b))" unfolding division_of_def by fastforce lemma division_of_subset: @@ -708,7 +779,7 @@ using assms(2) by auto show "k \ \q" using `k \ q` by auto - show "\a b. k = {a..b}" + show "\a b. k = cbox a b" using *(4)[OF kp] by auto show "k \ {}" using *(3)[OF kp] by auto @@ -725,7 +796,7 @@ unfolding division_of_def by auto lemma division_of_content_0: - assumes "content {a..b} = 0" "d division_of {a..b}" + assumes "content (cbox a b) = 0" "d division_of (cbox a b)" shows "\k\d. content k = 0" unfolding forall_in_division[OF assms(2)] apply (rule,rule,rule) @@ -737,7 +808,7 @@ qed lemma division_inter: - fixes s1 s2 :: "'a::ordered_euclidean_space set" + fixes s1 s2 :: "'a::euclidean_space set" assumes "p1 division_of s1" and "p2 division_of s2" shows "{k1 \ k2 | k1 k2 .k1 \ p1 \ k2 \ p2 \ k1 \ k2 \ {}} division_of (s1 \ s2)" @@ -771,11 +842,11 @@ show "k \ s1 \ s2" using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)] unfolding k by auto - obtain a1 b1 where k1: "k1 = {a1..b1}" + obtain a1 b1 where k1: "k1 = cbox a1 b1" using division_ofD(4)[OF assms(1) k(2)] by blast - obtain a2 b2 where k2: "k2 = {a2..b2}" + obtain a2 b2 where k2: "k2 = cbox a2 b2" using division_ofD(4)[OF assms(2) k(3)] by blast - show "\a b. k = {a..b}" + show "\a b. k = cbox a b" unfolding k k1 k2 unfolding inter_interval by auto } fix k1 k2 @@ -807,22 +878,22 @@ lemma division_inter_1: assumes "d division_of i" - and "{a..b::'a::ordered_euclidean_space} \ i" - shows "{{a..b} \ k | k. k \ d \ {a..b} \ k \ {}} division_of {a..b}" -proof (cases "{a..b} = {}") + and "cbox a (b::'a::euclidean_space) \ i" + shows "{cbox a b \ k | k. k \ d \ cbox a b \ k \ {}} division_of (cbox a b)" +proof (cases "cbox a b = {}") case True show ?thesis unfolding True and division_of_trivial by auto next case False - have *: "{a..b} \ i = {a..b}" using assms(2) by auto + have *: "cbox a b \ i = cbox a b" using assms(2) by auto show ?thesis using division_inter[OF division_of_self[OF False] assms(1)] unfolding * by auto qed lemma elementary_inter: - fixes s t :: "'a::ordered_euclidean_space set" + fixes s t :: "'a::euclidean_space set" assumes "p1 division_of s" and "p2 division_of t" shows "\p. p division_of (s \ t)" @@ -833,7 +904,7 @@ lemma elementary_inters: assumes "finite f" and "f \ {}" - and "\s\f. \p. p division_of (s::('a::ordered_euclidean_space) set)" + and "\s\f. \p. p division_of (s::('a::euclidean_space) set)" shows "\p. p division_of (\ f)" using assms proof (induct f rule: finite_induct) @@ -898,33 +969,32 @@ using k d1(2) d2(2) by auto show "k \ {}" using k d1(3) d2(3) by auto - show "\a b. k = {a..b}" + show "\a b. k = cbox a b" using k d1(4) d2(4) by auto qed lemma partial_division_extend_1: - fixes a b c d :: "'a::ordered_euclidean_space" - assumes incl: "{c..d} \ {a..b}" - and nonempty: "{c..d} \ {}" - obtains p where "p division_of {a..b}" "{c..d} \ p" + fixes a b c d :: "'a::euclidean_space" + assumes incl: "cbox c d \ cbox a b" + and nonempty: "cbox c d \ {}" + obtains p where "p division_of (cbox a b)" "cbox c d \ p" proof let ?B = "\f::'a\'a \ 'a. - {(\i\Basis. (fst (f i) \ i) *\<^sub>R i) .. (\i\Basis. (snd (f i) \ i) *\<^sub>R i)}" + cbox (\i\Basis. (fst (f i) \ i) *\<^sub>R i) (\i\Basis. (snd (f i) \ i) *\<^sub>R i)" def p \ "?B ` (Basis \\<^sub>E {(a, c), (c, d), (d, b)})" - show "{c .. d} \ p" + show "cbox c d \ p" unfolding p_def - by (auto simp add: interval_eq_empty eucl_le[where 'a='a] - intro!: image_eqI[where x="\(i::'a)\Basis. (c, d)"]) + by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\(i::'a)\Basis. (c, d)"]) { fix i :: 'a assume "i \ Basis" with incl nonempty have "a \ i \ c \ i" "c \ i \ d \ i" "d \ i \ b \ i" - unfolding interval_eq_empty subset_interval by (auto simp: not_le) + unfolding box_eq_empty subset_box by (auto simp: not_le) } note ord = this - show "p division_of {a..b}" + show "p division_of (cbox a b)" proof (rule division_ofI) show "finite p" unfolding p_def by (auto intro!: finite_PiE) @@ -933,10 +1003,10 @@ assume "k \ p" then obtain f where f: "f \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f" by (auto simp: p_def) - then show "\a b. k = {a..b}" + then show "\a b. k = cbox a b" by auto - have "k \ {a..b} \ k \ {}" - proof (simp add: k interval_eq_empty subset_interval not_less eucl_le[where 'a='a], safe) + have "k \ cbox a b \ k \ {}" + proof (simp add: k box_eq_empty subset_box not_less, safe) fix i :: 'a assume i: "i \ Basis" with f have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" @@ -945,7 +1015,7 @@ show "a \ i \ fst (f i) \ i" "snd (f i) \ i \ b \ i" "fst (f i) \ i \ snd (f i) \ i" by auto qed - then show "k \ {}" "k \ {a .. b}" + then show "k \ {}" "k \ cbox a b" by auto { fix l @@ -966,13 +1036,13 @@ "g i = (a, c) \ g i = (c, d) \ g i = (d, b)" using f g by (auto simp: PiE_iff) with * ord[of i] show "interior l \ interior k = {}" - by (auto simp add: l k interior_closed_interval disjoint_interval intro!: bexI[of _ i]) + by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i]) } - note `k \ {a.. b}` + note `k \ cbox a b` } moreover { - fix x assume x: "x \ {a .. b}" + fix x assume x: "x \ cbox a b" have "\i\Basis. \l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" proof fix i :: 'a @@ -980,7 +1050,7 @@ with x ord[of i] have "(a \ i \ x \ i \ x \ i \ c \ i) \ (c \ i \ x \ i \ x \ i \ d \ i) \ (d \ i \ x \ i \ x \ i \ b \ i)" - by (auto simp: eucl_le[where 'a='a]) + by (auto simp: cbox_def) then show "\l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" by auto qed @@ -990,21 +1060,21 @@ moreover from f have "restrict f Basis \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" by auto moreover from f have "x \ ?B (restrict f Basis)" - by (auto simp: mem_interval eucl_le[where 'a='a]) + by (auto simp: mem_box) ultimately have "\k\p. x \ k" unfolding p_def by blast } - ultimately show "\p = {a..b}" + ultimately show "\p = cbox a b" by auto qed qed lemma partial_division_extend_interval: - assumes "p division_of (\p)" "(\p) \ {a..b}" - obtains q where "p \ q" "q division_of {a..b::'a::ordered_euclidean_space}" + assumes "p division_of (\p)" "(\p) \ cbox a b" + obtains q where "p \ q" "q division_of cbox a (b::'a::euclidean_space)" proof (cases "p = {}") case True - obtain q where "q division_of {a..b}" + obtain q where "q division_of (cbox a b)" by (rule elementary_interval) then show ?thesis apply - @@ -1015,20 +1085,20 @@ next case False note p = division_ofD[OF assms(1)] - have *: "\k\p. \q. q division_of {a..b} \ k \ q" + have *: "\k\p. \q. q division_of cbox a b \ k \ q" proof case goal1 - obtain c d where k: "k = {c..d}" + obtain c d where k: "k = cbox c d" using p(4)[OF goal1] by blast - have *: "{c..d} \ {a..b}" "{c..d} \ {}" + have *: "cbox c d \ cbox a b" "cbox c d \ {}" using p(2,3)[OF goal1, unfolded k] using assms(2) by (blast intro: order.trans)+ - obtain q where "q division_of {a..b}" "{c..d} \ q" + obtain q where "q division_of cbox a b" "cbox c d \ q" by (rule partial_division_extend_1[OF *]) then show ?case unfolding k by auto qed - obtain q where q: "\x. x \ p \ q x division_of {a..b}" "\x. x \ p \ x \ q x" + obtain q where q: "\x. x \ p \ q x division_of cbox a b" "\x. x \ p \ x \ q x" using bchoice[OF *] by blast have "\x. x \ p \ \d. d division_of \(q x - {x})" apply rule @@ -1058,15 +1128,15 @@ apply (rule that[of "d \ p"]) proof - have *: "\s f t. s \ {} \ \i\s. f i \ i = t \ t = \(f ` s) \ \s" by auto - have *: "{a..b} = \((\i. \(q i - {i})) ` p) \ \p" + have *: "cbox a b = \((\i. \(q i - {i})) ` p) \ \p" apply (rule *[OF False]) proof fix i assume i: "i \ p" - show "\(q i - {i}) \ i = {a..b}" + show "\(q i - {i}) \ i = cbox a b" using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto qed - show "d \ p division_of {a..b}" + show "d \ p division_of (cbox a b)" unfolding * apply (rule division_disjoint_union[OF d assms(1)]) apply (rule inter_interior_unions_intervals) @@ -1084,7 +1154,7 @@ apply (rule inter_interior_unions_intervals) proof - note qk=division_ofD[OF q(1)[OF k]] - show "finite (q k - {k})" "open (interior k)" "\t\q k - {k}. \a b. t = {a..b}" + show "finite (q k - {k})" "open (interior k)" "\t\q k - {k}. \a b. t = cbox a b" using qk by auto show "\t\q k - {k}. interior k \ interior t = {}" using qk(5) using q(2)[OF k] by auto @@ -1101,19 +1171,19 @@ qed lemma elementary_bounded[dest]: - fixes s :: "'a::ordered_euclidean_space set" + fixes s :: "'a::euclidean_space set" shows "p division_of s \ bounded s" unfolding division_of_def by (metis bounded_Union bounded_interval) -lemma elementary_subset_interval: - "p division_of s \ \a b. s \ {a..b::'a::ordered_euclidean_space}" - by (meson elementary_bounded bounded_subset_closed_interval) +lemma elementary_subset_cbox: + "p division_of s \ \a b. s \ cbox a (b::'a::euclidean_space)" + by (meson elementary_bounded bounded_subset_cbox) lemma division_union_intervals_exists: - fixes a b :: "'a::ordered_euclidean_space" - assumes "{a..b} \ {}" - obtains p where "(insert {a..b} p) division_of ({a..b} \ {c..d})" -proof (cases "{c..d} = {}") + fixes a b :: "'a::euclidean_space" + assumes "cbox a b \ {}" + obtains p where "(insert (cbox a b) p) division_of (cbox a b \ cbox c d)" +proof (cases "cbox c d = {}") case True show ?thesis apply (rule that[of "{}"]) @@ -1124,29 +1194,29 @@ next case False show ?thesis - proof (cases "{a..b} \ {c..d} = {}") + proof (cases "cbox a b \ cbox c d = {}") case True have *: "\a b. {a, b} = {a} \ {b}" by auto show ?thesis - apply (rule that[of "{{c..d}}"]) + apply (rule that[of "{cbox c d}"]) unfolding * apply (rule division_disjoint_union) - using `{c..d} \ {}` True assms + using `cbox c d \ {}` True assms using interior_subset apply auto done next case False - obtain u v where uv: "{a..b} \ {c..d} = {u..v}" + obtain u v where uv: "cbox a b \ cbox c d = cbox u v" unfolding inter_interval by auto - have *: "{u..v} \ {c..d}" using uv by auto - obtain p where "p division_of {c..d}" "{u..v} \ p" + have *: "cbox u v \ cbox c d" using uv by auto + obtain p where "p division_of cbox c d" "cbox u v \ p" by (rule partial_division_extend_1[OF * False[unfolded uv]]) note p = this division_ofD[OF this(1)] - have *: "{a..b} \ {c..d} = {a..b} \ \(p - {{u..v}})" "\x s. insert x s = {x} \ s" + have *: "cbox a b \ cbox c d = cbox a b \ \(p - {cbox u v})" "\x s. insert x s = {x} \ s" using p(8) unfolding uv[symmetric] by auto show ?thesis - apply (rule that[of "p - {{u..v}}"]) + apply (rule that[of "p - {cbox u v}"]) unfolding *(1) apply (subst *(2)) apply (rule division_disjoint_union) @@ -1157,7 +1227,7 @@ unfolding interior_inter[symmetric] proof - have *: "\cd p uv ab. p \ cd \ ab \ cd = uv \ ab \ p = uv \ p" by auto - have "interior ({a..b} \ \(p - {{u..v}})) = interior({u..v} \ \(p - {{u..v}}))" + have "interior (cbox a b \ \(p - {cbox u v})) = interior(cbox u v \ \(p - {cbox u v}))" apply (rule arg_cong[of _ _ interior]) apply (rule *[OF _ uv]) using p(8) @@ -1169,7 +1239,7 @@ using p(6) p(7)[OF p(2)] p(3) apply auto done - finally show "interior ({a..b} \ \(p - {{u..v}})) = {}" . + finally show "interior (cbox a b \ \(p - {cbox u v})) = {}" . qed auto qed qed @@ -1191,9 +1261,9 @@ done lemma elementary_union_interval: - fixes a b :: "'a::ordered_euclidean_space" + fixes a b :: "'a::euclidean_space" assumes "p division_of \p" - obtains q where "q division_of ({a..b} \ \p)" + obtains q where "q division_of (cbox a b \ \p)" proof - note assm = division_ofD[OF assms] have lem1: "\f s. \\(f ` s) = \((\x. \(f x)) ` s)" @@ -1202,13 +1272,13 @@ by auto { presume "p = {} \ thesis" - "{a..b} = {} \ thesis" - "{a..b} \ {} \ interior {a..b} = {} \ thesis" - "p \ {} \ interior {a..b}\{} \ {a..b} \ {} \ thesis" + "cbox a b = {} \ thesis" + "cbox a b \ {} \ interior (cbox a b) = {} \ thesis" + "p \ {} \ interior (cbox a b)\{} \ cbox a b \ {} \ thesis" then show thesis by auto next assume as: "p = {}" - obtain p where "p division_of {a..b}" + obtain p where "p division_of (cbox a b)" by (rule elementary_interval) then show thesis apply - @@ -1217,7 +1287,7 @@ apply auto done next - assume as: "{a..b} = {}" + assume as: "cbox a b = {}" show thesis apply (rule that) unfolding as @@ -1225,9 +1295,9 @@ apply auto done next - assume as: "interior {a..b} = {}" "{a..b} \ {}" + assume as: "interior (cbox a b) = {}" "cbox a b \ {}" show thesis - apply (rule that[of "insert {a..b} p"],rule division_ofI) + apply (rule that[of "insert (cbox a b) p"],rule division_ofI) unfolding finite_insert apply (rule assm(1)) unfolding Union_insert using assm(2-4) as @@ -1235,25 +1305,25 @@ apply (fast dest: assm(5))+ done next - assume as: "p \ {}" "interior {a..b} \ {}" "{a..b} \ {}" - have "\k\p. \q. (insert {a..b} q) division_of ({a..b} \ k)" + assume as: "p \ {}" "interior (cbox a b) \ {}" "cbox a b \ {}" + have "\k\p. \q. (insert (cbox a b) q) division_of (cbox a b \ k)" proof case goal1 - from assm(4)[OF this] obtain c d where "k = {c..d}" by blast + from assm(4)[OF this] obtain c d where "k = cbox c d" by blast then show ?case apply - apply (rule division_union_intervals_exists[OF as(3), of c d]) apply auto done qed - from bchoice[OF this] obtain q where "\x\p. insert {a..b} (q x) division_of {a..b} \ x" .. + from bchoice[OF this] obtain q where "\x\p. insert (cbox a b) (q x) division_of (cbox a b) \ x" .. note q = division_ofD[OF this[rule_format]] - let ?D = "\{insert {a..b} (q k) | k. k \ p}" + let ?D = "\{insert (cbox a b) (q k) | k. k \ p}" show thesis apply (rule that[of "?D"]) apply (rule division_ofI) proof - - have *: "{insert {a..b} (q k) |k. k \ p} = (\k. insert {a..b} (q k)) ` p" + have *: "{insert (cbox a b) (q k) |k. k \ p} = (\k. insert (cbox a b) (q k)) ` p" by auto show "finite ?D" apply (rule finite_Union) @@ -1262,24 +1332,24 @@ using assm(1) q(1) apply auto done - show "\?D = {a..b} \ \p" + show "\?D = cbox a b \ \p" unfolding * lem1 - unfolding lem2[OF as(1), of "{a..b}", symmetric] + unfolding lem2[OF as(1), of "cbox a b", symmetric] using q(6) by auto fix k assume k: "k \ ?D" - then show "k \ {a..b} \ \p" + then show "k \ cbox a b \ \p" using q(2) by auto show "k \ {}" using q(3) k by auto - show "\a b. k = {a..b}" + show "\a b. k = cbox a b" using q(4) k by auto fix k' assume k': "k' \ ?D" "k \ k'" - obtain x where x: "k \ insert {a..b} (q x)" "x\p" + obtain x where x: "k \ insert (cbox a b) (q x)" "x\p" using k by auto - obtain x' where x': "k'\insert {a..b} (q x')" "x'\p" + obtain x' where x': "k'\insert (cbox a b) (q x')" "x'\p" using k' by auto show "interior k \ interior k' = {}" proof (cases "x = x'") @@ -1293,12 +1363,12 @@ next case False { - presume "k = {a..b} \ ?thesis" - and "k' = {a..b} \ ?thesis" - and "k \ {a..b} \ k' \ {a..b} \ ?thesis" + presume "k = cbox a b \ ?thesis" + and "k' = cbox a b \ ?thesis" + and "k \ cbox a b \ k' \ cbox a b \ ?thesis" then show ?thesis by auto next - assume as': "k = {a..b}" + assume as': "k = cbox a b" show ?thesis apply (rule q(5)) using x' k'(2) @@ -1306,7 +1376,7 @@ apply auto done next - assume as': "k' = {a..b}" + assume as': "k' = cbox a b" show ?thesis apply (rule q(5)) using x k'(2) @@ -1314,10 +1384,10 @@ apply auto done } - assume as': "k \ {a..b}" "k' \ {a..b}" - obtain c d where k: "k = {c..d}" + assume as': "k \ cbox a b" "k' \ cbox a b" + obtain c d where k: "k = cbox c d" using q(4)[OF x(2,1)] by blast - have "interior k \ interior {a..b} = {}" + have "interior k \ interior (cbox a b) = {}" apply (rule q(5)) using x k'(2) using as' @@ -1329,9 +1399,9 @@ apply auto done moreover - obtain c d where c_d: "k' = {c..d}" + obtain c d where c_d: "k' = cbox c d" using q(4)[OF x'(2,1)] by blast - have "interior k' \ interior {a..b} = {}" + have "interior k' \ interior (cbox a b) = {}" apply (rule q(5)) using x' k'(2) using as' @@ -1351,7 +1421,7 @@ lemma elementary_unions_intervals: assumes fin: "finite f" - and "\s. s \ f \ \a b. s = {a..b::'a::ordered_euclidean_space}" + and "\s. s \ f \ \a b. s = cbox a (b::'a::euclidean_space)" obtains p where "p division_of (\f)" proof - have "\p. p division_of (\f)" @@ -1361,7 +1431,7 @@ fix x F assume as: "finite F" "x \ F" "\p. p division_of \F" "x\f" from this(3) obtain p where p: "p division_of \F" .. - from assms(2)[OF as(4)] obtain a b where x: "x = {a..b}" by blast + from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast have *: "\F = \p" using division_ofD[OF p] by auto show "\p. p division_of \insert x F" @@ -1377,7 +1447,7 @@ qed lemma elementary_union: - fixes s t :: "'a::ordered_euclidean_space set" + fixes s t :: "'a::euclidean_space set" assumes "ps division_of s" and "pt division_of t" obtains p where "p division_of (s \ t)" @@ -1397,16 +1467,16 @@ qed lemma partial_division_extend: - fixes t :: "'a::ordered_euclidean_space set" + fixes t :: "'a::euclidean_space set" assumes "p division_of s" and "q division_of t" and "s \ t" obtains r where "p \ r" and "r division_of t" proof - note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)] - obtain a b where ab: "t \ {a..b}" - using elementary_subset_interval[OF assms(2)] by auto - obtain r1 where "p \ r1" "r1 division_of {a..b}" + obtain a b where ab: "t \ cbox a b" + using elementary_subset_cbox[OF assms(2)] by auto + obtain r1 where "p \ r1" "r1 division_of (cbox a b)" apply (rule partial_division_extend_interval) apply (rule assms(1)[unfolded divp(6)[symmetric]]) apply (rule subset_trans) @@ -1452,7 +1522,7 @@ proof - have "interior s \ interior (\(r1-p)) = {}" proof (rule inter_interior_unions_intervals) - show "finite (r1 - p)" and "open (interior s)" and "\t\r1-p. \a b. t = {a..b}" + show "finite (r1 - p)" and "open (interior s)" and "\t\r1-p. \a b. t = cbox a b" using r1 by auto have *: "\s. (\x. x \ s \ False) \ s = {}" by auto @@ -1462,7 +1532,7 @@ assume as: "m \ r1 - p" have "interior m \ interior (\p) = {}" proof (rule inter_interior_unions_intervals) - show "finite p" and "open (interior m)" and "\t\p. \a b. t = {a..b}" + show "finite p" and "open (interior m)" and "\t\p. \a b. t = cbox a b" using divp by auto show "\t\p. interior m \ interior t = {}" apply (rule, rule r1(7)) @@ -1486,7 +1556,7 @@ definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) where "s tagged_partial_division_of i \ finite s \ - (\x k. (x, k) \ s \ x \ k \ k \ i \ (\a b. k = {a..b})) \ + (\x k. (x, k) \ s \ x \ k \ k \ i \ (\a b. k = cbox a b)) \ (\x1 k1 x2 k2. (x1, k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ interior k1 \ interior k2 = {})" @@ -1495,7 +1565,7 @@ shows "finite s" and "\x k. (x,k) \ s \ x \ k" and "\x k. (x,k) \ s \ k \ i" - and "\x k. (x,k) \ s \ \a b. k = {a..b}" + and "\x k. (x,k) \ s \ \a b. k = cbox a b" and "\x1 k1 x2 k2. (x1,k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ interior k1 \ interior k2 = {}" using assms unfolding tagged_partial_division_of_def by blast+ @@ -1509,7 +1579,7 @@ lemma tagged_division_of: "s tagged_division_of i \ finite s \ - (\x k. (x, k) \ s \ x \ k \ k \ i \ (\a b. k = {a..b})) \ + (\x k. (x, k) \ s \ x \ k \ k \ i \ (\a b. k = cbox a b)) \ (\x1 k1 x2 k2. (x1, k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ interior k1 \ interior k2 = {}) \ (\{k. \x. (x,k) \ s} = i)" @@ -1519,7 +1589,7 @@ assumes "finite s" and "\x k. (x,k) \ s \ x \ k" and "\x k. (x,k) \ s \ k \ i" - and "\x k. (x,k) \ s \ \a b. k = {a..b}" + and "\x k. (x,k) \ s \ \a b. k = cbox a b" and "\x1 k1 x2 k2. (x1,k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ interior k1 \ interior k2 = {}" and "(\{k. \x. (x,k) \ s} = i)" @@ -1544,7 +1614,7 @@ shows "finite s" and "\x k. (x,k) \ s \ x \ k" and "\x k. (x,k) \ s \ k \ i" - and "\x k. (x,k) \ s \ \a b. k = {a..b}" + and "\x k. (x,k) \ s \ \a b. k = cbox a b" and "\x1 k1 x2 k2. (x1, k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ interior k1 \ interior k2 = {}" and "(\{k. \x. (x,k) \ s} = i)" @@ -1561,7 +1631,7 @@ assume k: "k \ snd ` s" then obtain xk where xk: "(xk, k) \ s" by auto - then show "k \ i" "k \ {}" "\a b. k = {a..b}" + then show "k \ i" "k \ {}" "\a b. k = cbox a b" using assm by fastforce+ fix k' assume k': "k' \ snd ` s" "k \ k'" @@ -1587,7 +1657,7 @@ assume k: "k \ snd ` s" then obtain xk where xk: "(xk, k) \ s" by auto - then show "k \ {}" "\a b. k = {a..b}" "k \ \(snd ` s)" + then show "k \ {}" "\a b. k = cbox a b" "k \ \(snd ` s)" using assm by auto fix k' assume k': "k' \ snd ` s" "k \ k'" @@ -1612,9 +1682,9 @@ by blast lemma setsum_over_tagged_division_lemma: - fixes d :: "'m::ordered_euclidean_space set \ 'a::real_normed_vector" + fixes d :: "'m::euclidean_space set \ 'a::real_normed_vector" assumes "p tagged_division_of i" - and "\u v. {u..v} \ {} \ content {u..v} = 0 \ d {u..v} = 0" + and "\u v. cbox u v \ {} \ content (cbox u v) = 0 \ d (cbox u v) = 0" shows "setsum (\(x,k). d k) p = setsum d (snd ` p)" proof - note assm = tagged_division_ofD[OF assms(1)] @@ -1628,7 +1698,7 @@ using assm by auto fix x y assume as: "x\p" "y\p" "x\y" "snd x = snd y" - obtain a b where ab: "snd x = {a..b}" + obtain a b where ab: "snd x = cbox a b" using assm(4)[of "fst x" "snd x"] as(1) by auto have "(fst x, snd y) \ p" "(fst x, snd y) \ y" unfolding as(4)[symmetric] using as(1-3) by auto @@ -1638,9 +1708,9 @@ using as apply auto done - then have "content {a..b} = 0" + then have "content (cbox a b) = 0" unfolding as(4)[symmetric] ab content_eq_0_interior by auto - then have "d {a..b} = 0" + then have "d (cbox a b) = 0" apply - apply (rule assms(2)) using assm(2)[of "fst x" "snd x"] as(1) @@ -1664,9 +1734,13 @@ lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \ p = {}" unfolding tagged_division_of by auto -lemma tagged_division_of_self: "x \ {a..b} \ {(x,{a..b})} tagged_division_of {a..b}" +lemma tagged_division_of_self: "x \ cbox a b \ {(x,cbox a b)} tagged_division_of (cbox a b)" by (rule tagged_division_ofI) auto +lemma tagged_division_of_self_real: "x \ {a .. b::real} \ {(x,{a .. b})} tagged_division_of {a .. b}" + unfolding box_real[symmetric] + by (rule tagged_division_of_self) + lemma tagged_division_union: assumes "p1 tagged_division_of s1" and "p2 tagged_division_of s2" @@ -1681,7 +1755,7 @@ using p1(6) p2(6) by blast fix x k assume xk: "(x, k) \ p1 \ p2" - show "x \ k" "\a b. k = {a..b}" + show "x \ k" "\a b. k = cbox a b" using xk p1(2,4) p2(2,4) by auto show "k \ s1 \ s2" using xk p1(3) p2(3) by blast @@ -1727,7 +1801,7 @@ assume xk: "(x, k) \ \(pfn ` iset)" then obtain i where i: "i \ iset" "(x, k) \ pfn i" by auto - show "x \ k" "\a b. k = {a..b}" "k \ \iset" + show "x \ k" "\a b. k = cbox a b" "k \ \iset" using assm(2-4)[OF i] using i(1) by auto fix x' k' assume xk': "(x', k') \ \(pfn ` iset)" "(x, k) \ (x', k')" @@ -1806,48 +1880,56 @@ norm (setsum (\(x,k). content k *\<^sub>R f x) p - y) < e))" definition has_integral :: - "('n::ordered_euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" + "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" (infixr "has'_integral" 46) where "(f has_integral y) i \ - (if \a b. i = {a..b} + (if \a b. i = cbox a b then (f has_integral_compact_interval y) i - else (\e>0. \B>0. \a b. ball 0 B \ {a..b} \ - (\z. ((\x. if x \ i then f x else 0) has_integral_compact_interval z) {a..b} \ + else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ i then f x else 0) has_integral_compact_interval z) (cbox a b) \ norm (z - y) < e)))" lemma has_integral: - "(f has_integral y) {a..b} \ + "(f has_integral y) (cbox a b) \ (\e>0. \d. gauge d \ - (\p. p tagged_division_of {a..b} \ d fine p \ + (\p. p tagged_division_of (cbox a b) \ d fine p \ norm (setsum (\(x,k). content(k) *\<^sub>R f x) p - y) < e))" unfolding has_integral_def has_integral_compact_interval_def by auto +lemma has_integral_real: + "(f has_integral y) {a .. b::real} \ + (\e>0. \d. gauge d \ + (\p. p tagged_division_of {a .. b} \ d fine p \ + norm (setsum (\(x,k). content(k) *\<^sub>R f x) p - y) < e))" + unfolding box_real[symmetric] + by (rule has_integral) + lemma has_integralD[dest]: - assumes "(f has_integral y) ({a..b})" + assumes "(f has_integral y) (cbox a b)" and "e > 0" obtains d where "gauge d" - and "\p. p tagged_division_of {a..b} \ d fine p \ + and "\p. p tagged_division_of (cbox a b) \ d fine p \ norm (setsum (\(x,k). content(k) *\<^sub>R f(x)) p - y) < e" using assms unfolding has_integral by auto lemma has_integral_alt: "(f has_integral y) i \ - (if \a b. i = {a..b} + (if \a b. i = cbox a b then (f has_integral y) i - else (\e>0. \B>0. \a b. ball 0 B \ {a..b} \ - (\z. ((\x. if x \ i then f(x) else 0) has_integral z) ({a..b}) \ norm (z - y) < e)))" + else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm (z - y) < e)))" unfolding has_integral unfolding has_integral_compact_interval_def has_integral_def by auto lemma has_integral_altD: assumes "(f has_integral y) i" - and "\ (\a b. i = {a..b})" + and "\ (\a b. i = cbox a b)" and "e>0" obtains B where "B > 0" - and "\a b. ball 0 B \ {a..b} \ - (\z. ((\x. if x \ i then f(x) else 0) has_integral z) ({a..b}) \ norm(z - y) < e)" + and "\a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm(z - y) < e)" using assms unfolding has_integral unfolding has_integral_compact_interval_def has_integral_def @@ -1868,8 +1950,8 @@ by auto lemma setsum_content_null: - assumes "content {a..b} = 0" - and "p tagged_division_of {a..b}" + assumes "content (cbox a b) = 0" + and "p tagged_division_of (cbox a b)" shows "setsum (\(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)" proof (rule setsum_0', rule) fix y @@ -1877,7 +1959,7 @@ obtain x k where xk: "y = (x, k)" using surj_pair[of y] by blast note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]] - from this(2) obtain c d where k: "k = {c..d}" by blast + from this(2) obtain c d where k: "k = cbox c d" by blast have "(\(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x" unfolding xk by auto also have "\ = 0" @@ -1923,22 +2005,22 @@ subsection {* General bisection principle for intervals; might be useful elsewhere. *} lemma interval_bisection_step: - fixes type :: "'a::ordered_euclidean_space" + fixes type :: "'a::euclidean_space" assumes "P {}" and "\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P (s \ t)" - and "\ P {a..b::'a}" - obtains c d where "\ P{c..d}" + and "\ P (cbox a (b::'a))" + obtains c d where "\ P (cbox c d)" and "\i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" proof - - have "{a..b} \ {}" + have "cbox a b \ {}" using assms(1,3) by metis then have ab: "\i. i\Basis \ a \ i \ b \ i" - by (auto simp: eucl_le[where 'a='a]) + by (force simp: mem_box) { fix f have "finite f \ \s\f. P s \ - \s\f. \a b. s = {a..b} \ + \s\f. \a b. s = cbox a b \ \s\f.\t\f. s \ t \ interior s \ interior t = {} \ P (\f)" proof (induct f rule: finite_induct) case empty @@ -1959,11 +2041,11 @@ done qed } note * = this - let ?A = "{{c..d} | c d::'a. \i\Basis. (c\i = a\i) \ (d\i = (a\i + b\i) / 2) \ + let ?A = "{cbox c d | c d::'a. \i\Basis. (c\i = a\i) \ (d\i = (a\i + b\i) / 2) \ (c\i = (a\i + b\i) / 2) \ (d\i = b\i)}" let ?PP = "\c d. \i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" { - presume "\c d. ?PP c d \ P {c..d} \ False" + presume "\c d. ?PP c d \ P (cbox c d) \ False" then show thesis unfolding atomize_not not_all apply - @@ -1972,23 +2054,23 @@ apply auto done } - assume as: "\c d. ?PP c d \ P {c..d}" + assume as: "\c d. ?PP c d \ P (cbox c d)" have "P (\ ?A)" apply (rule *) apply (rule_tac[2-] ballI) apply (rule_tac[4] ballI) apply (rule_tac[4] impI) proof - - let ?B = "(\s.{(\i\Basis. (if i \ s then a\i else (a\i + b\i) / 2) *\<^sub>R i)::'a .. - (\i\Basis. (if i \ s then (a\i + b\i) / 2 else b\i) *\<^sub>R i)}) ` {s. s \ Basis}" + let ?B = "(\s. cbox (\i\Basis. (if i \ s then a\i else (a\i + b\i) / 2) *\<^sub>R i::'a) + (\i\Basis. (if i \ s then (a\i + b\i) / 2 else b\i) *\<^sub>R i)) ` {s. s \ Basis}" have "?A \ ?B" proof case goal1 - then obtain c d where x: "x = {c..d}" + then obtain c d where x: "x = cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" by blast - have *: "\a b c d. a = c \ b = d \ {a..b} = {c..d}" + have *: "\a b c d. a = c \ b = d \ cbox a b = cbox c d" by auto show "x \ ?B" unfolding image_iff @@ -2011,7 +2093,7 @@ fix s assume "s \ ?A" then obtain c d where s: - "s = {c..d}" + "s = cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" @@ -2024,12 +2106,12 @@ then show ?case using s(2)[of i] using ab[OF `i \ Basis`] by auto qed - show "\a b. s = {a..b}" + show "\a b. s = cbox a b" unfolding s by auto fix t assume "t \ ?A" then obtain e f where t: - "t = {e..f}" + "t = cbox e f" "\i. i \ Basis \ e \ i = a \ i \ f \ i = (a \ i + b \ i) / 2 \ e \ i = (a \ i + b \ i) / 2 \ f \ i = b \ i" @@ -2054,12 +2136,12 @@ have *: "\s t. (\a. a \ s \ a \ t \ False) \ s \ t = {}" by auto show "interior s \ interior t = {}" - unfolding s t interior_closed_interval + unfolding s t interior_cbox proof (rule *) fix x assume "x \ box c d" "x \ box e f" then have x: "c\i < d\i" "e\i < f\i" "c\i < f\i" "e\i < d\i" - unfolding mem_interval using i' + unfolding mem_box using i' apply - apply (erule_tac[!] x=i in ballE)+ apply auto @@ -2080,35 +2162,35 @@ qed qed qed - also have "\ ?A = {a..b}" + also have "\ ?A = cbox a b" proof (rule set_eqI,rule) fix x assume "x \ \?A" then obtain c d where x: - "x \ {c..d}" + "x \ cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" by blast - show "x\{a..b}" - unfolding mem_interval + show "x\cbox a b" + unfolding mem_box proof safe fix i :: 'a assume i: "i \ Basis" then show "a \ i \ x \ i" "x \ i \ b \ i" - using x(2)[OF i] x(1)[unfolded mem_interval,THEN bspec, OF i] by auto + using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto qed next fix x - assume x: "x \ {a..b}" + assume x: "x \ cbox a b" have "\i\Basis. \c d. (c = a\i \ d = (a\i + b\i) / 2 \ c = (a\i + b\i) / 2 \ d = b\i) \ c\x\i \ x\i \ d" (is "\i\Basis. \c d. ?P i c d") - unfolding mem_interval + unfolding mem_box proof fix i :: 'a assume i: "i \ Basis" have "?P i (a\i) ((a \ i + b \ i) / 2) \ ?P i ((a \ i + b \ i) / 2) (b\i)" - using x[unfolded mem_interval,THEN bspec, OF i] by auto + using x[unfolded mem_box,THEN bspec, OF i] by auto then show "\c d. ?P i c d" by blast qed @@ -2116,8 +2198,8 @@ unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff apply - apply (erule exE)+ - apply (rule_tac x="{xa..xaa}" in exI) - unfolding mem_interval + apply (rule_tac x="cbox xa xaa" in exI) + unfolding mem_box apply auto done qed @@ -2126,25 +2208,25 @@ qed lemma interval_bisection: - fixes type :: "'a::ordered_euclidean_space" + fixes type :: "'a::euclidean_space" assumes "P {}" and "(\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P(s \ t))" - and "\ P {a..b::'a}" - obtains x where "x \ {a..b}" - and "\e>0. \c d. x \ {c..d} \ {c..d} \ ball x e \ {c..d} \ {a..b} \ \ P {c..d}" -proof - - have "\x. \y. \ P {fst x..snd x} \ (\ P {fst y..snd y} \ + and "\ P (cbox a (b::'a))" + obtains x where "x \ cbox a b" + and "\e>0. \c d. x \ cbox c d \ cbox c d \ ball x e \ cbox c d \ cbox a b \ \ P (cbox c d)" +proof - + have "\x. \y. \ P (cbox (fst x) (snd x)) \ (\ P (cbox (fst y) (snd y)) \ (\i\Basis. fst x\i \ fst y\i \ fst y\i \ snd y\i \ snd y\i \ snd x\i \ 2 * (snd y\i - fst y\i) \ snd x\i - fst x\i))" proof case goal1 then show ?case proof - - presume "\ P {fst x..snd x} \ ?thesis" - then show ?thesis by (cases "P {fst x..snd x}") auto + presume "\ P (cbox (fst x) (snd x)) \ ?thesis" + then show ?thesis by (cases "P (cbox (fst x) (snd x))") auto next - assume as: "\ P {fst x..snd x}" - obtain c d where "\ P {c..d}" + assume as: "\ P (cbox (fst x) (snd x))" + obtain c d where "\ P (cbox c d)" "\i\Basis. fst x \ i \ c \ i \ c \ i \ d \ i \ @@ -2160,8 +2242,8 @@ qed then obtain f where f: "\x. - \ P {fst x..snd x} \ - \ P {fst (f x)..snd (f x)} \ + \ P (cbox (fst x) (snd x)) \ + \ P (cbox (fst (f x)) (snd (f x))) \ (\i\Basis. fst x \ i \ fst (f x) \ i \ fst (f x) \ i \ snd (f x) \ i \ @@ -2175,7 +2257,7 @@ def A \ "\n. fst(AB n)" def B \ "\n. snd(AB n)" note ab_def = A_def B_def AB_def - have "A 0 = a" "B 0 = b" "\n. \ P {A(Suc n)..B(Suc n)} \ + have "A 0 = a" "B 0 = b" "\n. \ P (cbox (A(Suc n)) (B(Suc n))) \ (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ 2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i)" (is "\n. ?P n") proof - @@ -2204,7 +2286,7 @@ qed note AB = this(1-2) conjunctD2[OF this(3),rule_format] - have interv: "\e. 0 < e \ \n. \x\{A n..B n}. \y\{A n..B n}. dist x y < e" + have interv: "\e. 0 < e \ \n. \x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" proof - case goal1 obtain n where n: "(\i\Basis. b \ i - a \ i) / e < 2 ^ n" @@ -2215,7 +2297,7 @@ apply rule proof - fix x y - assume xy: "x\{A n..B n}" "y\{A n..B n}" + assume xy: "x\cbox (A n) (B n)" "y\cbox (A n) (B n)" have "dist x y \ setsum (\i. abs((x - y)\i)) Basis" unfolding dist_norm by(rule norm_le_l1) also have "\ \ setsum (\i. B n\i - A n\i) Basis" @@ -2223,7 +2305,7 @@ fix i :: 'a assume i: "i \ Basis" show "\(x - y) \ i\ \ B n \ i - A n \ i" - using xy[unfolded mem_interval,THEN bspec, OF i] + using xy[unfolded mem_box,THEN bspec, OF i] by (auto simp: inner_diff_left) qed also have "\ \ setsum (\i. b\i - a\i) Basis / 2^n" @@ -2251,27 +2333,27 @@ qed { fix n m :: nat - assume "m \ n" then have "{A n..B n} \ {A m..B m}" + assume "m \ n" then have "cbox (A n) (B n) \ cbox (A m) (B m)" proof (induction rule: inc_induct) case (step i) show ?case - using AB(4) by (intro order_trans[OF step.IH] subset_interval_imp) auto + using AB(4) by (intro order_trans[OF step.IH] subset_box_imp) auto qed simp } note ABsubset = this - have "\a. \n. a\{A n..B n}" - by (rule decreasing_closed_nest[rule_format,OF closed_interval _ ABsubset interv]) + have "\a. \n. a\ cbox (A n) (B n)" + by (rule decreasing_closed_nest[rule_format,OF closed_cbox _ ABsubset interv]) (metis nat.exhaust AB(1-3) assms(1,3)) - then obtain x0 where x0: "\n. x0 \ {A n..B n}" + then obtain x0 where x0: "\n. x0 \ cbox (A n) (B n)" by blast show thesis proof (rule that[rule_format, of x0]) - show "x0\{a..b}" + show "x0\cbox a b" using x0[of 0] unfolding AB . fix e :: real assume "e > 0" from interv[OF this] obtain n - where n: "\x\{A n..B n}. \y\{A n..B n}. dist x y < e" .. - show "\c d. x0 \ {c..d} \ {c..d} \ ball x0 e \ {c..d} \ {a..b} \ \ P {c..d}" + where n: "\x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" .. + show "\c d. x0 \ cbox c d \ cbox c d \ ball x0 e \ cbox c d \ cbox a b \ \ P (cbox c d)" apply (rule_tac x="A n" in exI) apply (rule_tac x="B n" in exI) apply rule @@ -2280,14 +2362,14 @@ defer apply rule proof - - show "\ P {A n..B n}" + show "\ P (cbox (A n) (B n))" apply (cases "0 < n") using AB(3)[of "n - 1"] assms(3) AB(1-2) apply auto done - show "{A n..B n} \ ball x0 e" + show "cbox (A n) (B n) \ ball x0 e" using n using x0[of n] by auto - show "{A n..B n} \ {a..b}" + show "cbox (A n) (B n) \ cbox a b" unfolding AB(1-2)[symmetric] by (rule ABsubset) auto qed qed @@ -2297,24 +2379,24 @@ subsection {* Cousin's lemma. *} lemma fine_division_exists: - fixes a b :: "'a::ordered_euclidean_space" + fixes a b :: "'a::euclidean_space" assumes "gauge g" - obtains p where "p tagged_division_of {a..b}" "g fine p" -proof - - presume "\ (\p. p tagged_division_of {a..b} \ g fine p) \ False" - then obtain p where "p tagged_division_of {a..b}" "g fine p" + obtains p where "p tagged_division_of (cbox a b)" "g fine p" +proof - + presume "\ (\p. p tagged_division_of (cbox a b) \ g fine p) \ False" + then obtain p where "p tagged_division_of (cbox a b)" "g fine p" by blast then show thesis .. next - assume as: "\ (\p. p tagged_division_of {a..b} \ g fine p)" + assume as: "\ (\p. p tagged_division_of (cbox a b) \ g fine p)" obtain x where x: - "x \ {a..b}" + "x \ (cbox a b)" "\e. 0 < e \ \c d. - x \ {c..d} \ - {c..d} \ ball x e \ - {c..d} \ {a..b} \ - \ (\p. p tagged_division_of {c..d} \ g fine p)" + x \ cbox c d \ + cbox c d \ ball x e \ + cbox c d \ (cbox a b) \ + \ (\p. p tagged_division_of cbox c d \ g fine p)" apply (rule interval_bisection[of "\s. \p. p tagged_division_of s \ g fine p",rule_format,OF _ _ as]) apply (rule_tac x="{}" in exI) defer @@ -2338,22 +2420,27 @@ obtain e where e: "e > 0" "ball x e \ g x" using gaugeD[OF assms, of x] unfolding open_contains_ball by auto from x(2)[OF e(1)] obtain c d where c_d: - "x \ {c..d}" - "{c..d} \ ball x e" - "{c..d} \ {a..b}" - "\ (\p. p tagged_division_of {c..d} \ g fine p)" + "x \ cbox c d" + "cbox c d \ ball x e" + "cbox c d \ cbox a b" + "\ (\p. p tagged_division_of cbox c d \ g fine p)" by blast - have "g fine {(x, {c..d})}" + have "g fine {(x, cbox c d)}" unfolding fine_def using e using c_d(2) by auto then show False using tagged_division_of_self[OF c_d(1)] using c_d by auto qed +lemma fine_division_exists_real: + fixes a b :: real + assumes "gauge g" + obtains p where "p tagged_division_of {a .. b}" "g fine p" + by (metis assms box_real(2) fine_division_exists) subsection {* Basic theorems about integrals. *} lemma has_integral_unique: - fixes f :: "'n::ordered_euclidean_space \ 'a::real_normed_vector" + fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k1) i" and "(f has_integral k2) i" shows "k1 = k2" @@ -2363,23 +2450,23 @@ then have e: "?e > 0" by auto have lem: "\f::'n \ 'a. \a b k1 k2. - (f has_integral k1) ({a..b}) \ (f has_integral k2) ({a..b}) \ k1 \ k2 \ False" + (f has_integral k1) (cbox a b) \ (f has_integral k2) (cbox a b) \ k1 \ k2 \ False" proof - case goal1 let ?e = "norm (k1 - k2) / 2" from goal1(3) have e: "?e > 0" by auto obtain d1 where d1: "gauge d1" - "\p. p tagged_division_of {a..b} \ + "\p. p tagged_division_of cbox a b \ d1 fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - k1) < norm (k1 - k2) / 2" by (rule has_integralD[OF goal1(1) e]) blast obtain d2 where d2: "gauge d2" - "\p. p tagged_division_of {a..b} \ + "\p. p tagged_division_of cbox a b \ d2 fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - k2) < norm (k1 - k2) / 2" by (rule has_integralD[OF goal1(2) e]) blast obtain p where p: - "p tagged_division_of {a..b}" + "p tagged_division_of cbox a b" "(\x. d1 x \ d2 x) fine p" by (rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)]]) let ?c = "(\(x, k)\p. content k *\<^sub>R f x)" @@ -2395,40 +2482,40 @@ finally show False by auto qed { - presume "\ (\a b. i = {a..b}) \ False" + presume "\ (\a b. i = cbox a b) \ False" then show False apply - - apply (cases "\a b. i = {a..b}") + apply (cases "\a b. i = cbox a b") using assms apply (auto simp add:has_integral intro:lem[OF _ _ as]) done } - assume as: "\ (\a b. i = {a..b})" + assume as: "\ (\a b. i = cbox a b)" obtain B1 where B1: "0 < B1" - "\a b. ball 0 B1 \ {a..b} \ - \z. ((\x. if x \ i then f x else 0) has_integral z) {a..b} \ + "\a b. ball 0 B1 \ cbox a b \ + \z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ norm (z - k1) < norm (k1 - k2) / 2" by (rule has_integral_altD[OF assms(1) as,OF e]) blast obtain B2 where B2: "0 < B2" - "\a b. ball 0 B2 \ {a..b} \ - \z. ((\x. if x \ i then f x else 0) has_integral z) {a..b} \ + "\a b. ball 0 B2 \ cbox a b \ + \z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ norm (z - k2) < norm (k1 - k2) / 2" by (rule has_integral_altD[OF assms(2) as,OF e]) blast - have "\a b::'n. ball 0 B1 \ ball 0 B2 \ {a..b}" - apply (rule bounded_subset_closed_interval) + have "\a b::'n. ball 0 B1 \ ball 0 B2 \ cbox a b" + apply (rule bounded_subset_cbox) using bounded_Un bounded_ball apply auto done - then obtain a b :: 'n where ab: "ball 0 B1 \ {a..b}" "ball 0 B2 \ {a..b}" + then obtain a b :: 'n where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by blast obtain w where w: - "((\x. if x \ i then f x else 0) has_integral w) {a..b}" + "((\x. if x \ i then f x else 0) has_integral w) (cbox a b)" "norm (w - k1) < norm (k1 - k2) / 2" using B1(2)[OF ab(1)] by blast obtain z where z: - "((\x. if x \ i then f x else 0) has_integral z) {a..b}" + "((\x. if x \ i then f x else 0) has_integral z) (cbox a b)" "norm (z - k2) < norm (k1 - k2) / 2" using B2(2)[OF ab(2)] by blast have "z = w" @@ -2448,21 +2535,21 @@ by (rule some_equality) (auto intro: has_integral_unique) lemma has_integral_is_0: - fixes f :: "'n::ordered_euclidean_space \ 'a::real_normed_vector" + fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x\s. f x = 0" shows "(f has_integral 0) s" proof - have lem: "\a b. \f::'n \ 'a. - (\x\{a..b}. f(x) = 0) \ (f has_integral 0) ({a..b})" + (\x\cbox a b. f(x) = 0) \ (f has_integral 0) (cbox a b)" unfolding has_integral apply rule apply rule proof - fix a b e fix f :: "'n \ 'a" - assume as: "\x\{a..b}. f x = 0" "0 < (e::real)" + assume as: "\x\cbox a b. f x = 0" "0 < (e::real)" show "\d. gauge d \ - (\p. p tagged_division_of {a..b} \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e)" + (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e)" apply (rule_tac x="\x. ball x 1" in exI) apply rule apply (rule gaugeI) @@ -2491,10 +2578,10 @@ qed auto qed { - presume "\ (\a b. s = {a..b}) \ ?thesis" + presume "\ (\a b. s = cbox a b) \ ?thesis" then show ?thesis apply - - apply (cases "\a b. s = {a..b}") + apply (cases "\a b. s = cbox a b") using assms apply (auto simp add:has_integral intro: lem) done @@ -2504,7 +2591,7 @@ using assms apply auto done - assume "\ (\a b. s = {a..b})" + assume "\ (\a b. s = cbox a b)" then show ?thesis apply (subst has_integral_alt) unfolding if_not_P * @@ -2520,7 +2607,7 @@ fix e :: real fix a b assume "e > 0" - then show "\z. ((\x::'n. 0::'a) has_integral z) {a..b} \ norm (z - 0) < e" + then show "\z. ((\x::'n. 0::'a) has_integral z) (cbox a b) \ norm (z - 0) < e" apply (rule_tac x=0 in exI) apply(rule,rule lem) apply auto @@ -2528,14 +2615,14 @@ qed auto qed -lemma has_integral_0[simp]: "((\x::'n::ordered_euclidean_space. 0) has_integral 0) s" +lemma has_integral_0[simp]: "((\x::'n::euclidean_space. 0) has_integral 0) s" by (rule has_integral_is_0) auto lemma has_integral_0_eq[simp]: "((\x. 0) has_integral i) s \ i = 0" using has_integral_unique[OF has_integral_0] by auto lemma has_integral_linear: - fixes f :: "'n::ordered_euclidean_space \ 'a::real_normed_vector" + fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral y) s" and "bounded_linear h" shows "((h o f) has_integral ((h y))) s" @@ -2545,7 +2632,7 @@ from pos_bounded obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" by blast have lem: "\(f :: 'n \ 'a) y a b. - (f has_integral y) {a..b} \ ((h o f) has_integral h y) {a..b}" + (f has_integral y) (cbox a b) \ ((h o f) has_integral h y) (cbox a b)" apply (subst has_integral) apply rule apply rule @@ -2561,7 +2648,7 @@ done obtain g where g: "gauge g" - "\p. p tagged_division_of {a..b} \ g fine p \ + "\p. p tagged_division_of (cbox a b) \ g fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - y) < e / B" by (rule has_integralD[OF goal1(1) *]) blast show ?case @@ -2573,7 +2660,7 @@ apply (erule conjE) proof - fix p - assume as: "p tagged_division_of {a..b}" "g fine p" + assume as: "p tagged_division_of (cbox a b)" "g fine p" have *: "\x k. h ((\(x, k). content k *\<^sub>R f x) x) = (\(x, k). h (content k *\<^sub>R f x)) x" by auto have "(\(x, k)\p. content k *\<^sub>R (h \ f) x) = setsum (h \ (\(x, k). content k *\<^sub>R f x)) p" @@ -2590,15 +2677,15 @@ qed qed { - presume "\ (\a b. s = {a..b}) \ ?thesis" + presume "\ (\a b. s = cbox a b) \ ?thesis" then show ?thesis apply - - apply (cases "\a b. s = {a..b}") + apply (cases "\a b. s = cbox a b") using assms apply (auto simp add:has_integral intro!:lem) done } - assume as: "\ (\a b. s = {a..b})" + assume as: "\ (\a b. s = cbox a b)" then show ?thesis apply (subst has_integral_alt) unfolding if_not_P @@ -2611,11 +2698,11 @@ by (rule divide_pos_pos,rule e,rule B(1)) obtain M where M: "M > 0" - "\a b. ball 0 M \ {a..b} \ - \z. ((\x. if x \ s then f x else 0) has_integral z) {a..b} \ norm (z - y) < e / B" + "\a b. ball 0 M \ cbox a b \ + \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e / B" using has_integral_altD[OF assms(1) as *] by blast - show "\B>0. \a b. ball 0 B \ {a..b} \ - (\z. ((\x. if x \ s then (h \ f) x else 0) has_integral z) {a..b} \ norm (z - h y) < e)" + show "\B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ s then (h \ f) x else 0) has_integral z) (cbox a b) \ norm (z - h y) < e)" apply (rule_tac x=M in exI) apply rule apply (rule M(1)) @@ -2625,7 +2712,7 @@ proof - case goal1 obtain z where z: - "((\x. if x \ s then f x else 0) has_integral z) {a..b}" + "((\x. if x \ s then f x else 0) has_integral z) (cbox a b)" "norm (z - y) < e / B" using M(2)[OF goal1(1)] by blast have *: "(\x. if x \ s then (h \ f) x else 0) = h \ (\x. if x \ s then f x else 0)" @@ -2673,15 +2760,15 @@ done lemma has_integral_add: - fixes f :: "'n::ordered_euclidean_space \ 'a::real_normed_vector" + fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k) s" and "(g has_integral l) s" shows "((\x. f x + g x) has_integral (k + l)) s" proof - have lem:"\(f:: 'n \ 'a) g a b k l. - (f has_integral k) {a..b} \ - (g has_integral l) {a..b} \ - ((\x. f x + g x) has_integral (k + l)) {a..b}" + (f has_integral k) (cbox a b) \ + (g has_integral l) (cbox a b) \ + ((\x. f x + g x) has_integral (k + l)) (cbox a b)" proof - case goal1 show ?case @@ -2695,15 +2782,15 @@ by auto obtain d1 where d1: "gauge d1" - "\p. p tagged_division_of {a..b} \ d1 fine p \ + "\p. p tagged_division_of (cbox a b) \ d1 fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - k) < e / 2" using has_integralD[OF goal1(1) *] by blast obtain d2 where d2: "gauge d2" - "\p. p tagged_division_of {a..b} \ d2 fine p \ + "\p. p tagged_division_of (cbox a b) \ d2 fine p \ norm ((\(x, k)\p. content k *\<^sub>R g x) - l) < e / 2" using has_integralD[OF goal1(2) *] by blast - show "\d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ + show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R (f x + g x)) - (k + l)) < e)" apply (rule_tac x="\x. (d1 x) \ (d2 x)" in exI) apply rule @@ -2711,7 +2798,7 @@ apply (rule,rule,erule conjE) proof - fix p - assume as: "p tagged_division_of {a..b}" "(\x. d1 x \ d2 x) fine p" + assume as: "p tagged_division_of (cbox a b)" "(\x. d1 x \ d2 x) fine p" have *: "(\(x, k)\p. content k *\<^sub>R (f x + g x)) = (\(x, k)\p. content k *\<^sub>R f x) + (\(x, k)\p. content k *\<^sub>R g x)" unfolding scaleR_right_distrib setsum_addf[of "\(x,k). content k *\<^sub>R f x" "\(x,k). content k *\<^sub>R g x" p,symmetric] @@ -2735,15 +2822,15 @@ qed qed { - presume "\ (\a b. s = {a..b}) \ ?thesis" + presume "\ (\a b. s = cbox a b) \ ?thesis" then show ?thesis apply - - apply (cases "\a b. s = {a..b}") + apply (cases "\a b. s = cbox a b") using assms apply (auto simp add:has_integral intro!:lem) done } - assume as: "\ (\a b. s = {a..b})" + assume as: "\ (\a b. s = cbox a b)" then show ?thesis apply (subst has_integral_alt) unfolding if_not_P @@ -2756,14 +2843,14 @@ from has_integral_altD[OF assms(1) as *] obtain B1 where B1: "0 < B1" - "\a b. ball 0 B1 \ {a..b} \ - \z. ((\x. if x \ s then f x else 0) has_integral z) {a..b} \ norm (z - k) < e / 2" + "\a b. ball 0 B1 \ cbox a b \ + \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - k) < e / 2" by blast from has_integral_altD[OF assms(2) as *] obtain B2 where B2: "0 < B2" - "\a b. ball 0 B2 \ {a..b} \ - \z. ((\x. if x \ s then g x else 0) has_integral z) {a..b} \ norm (z - l) < e / 2" + "\a b. ball 0 B2 \ (cbox a b) \ + \z. ((\x. if x \ s then g x else 0) has_integral z) (cbox a b) \ norm (z - l) < e / 2" by blast show ?case apply (rule_tac x="max B1 B2" in exI) @@ -2775,21 +2862,21 @@ apply rule proof - fix a b - assume "ball 0 (max B1 B2) \ {a..b::'n}" - then have *: "ball 0 B1 \ {a..b::'n}" "ball 0 B2 \ {a..b::'n}" + assume "ball 0 (max B1 B2) \ cbox a (b::'n)" + then have *: "ball 0 B1 \ cbox a (b::'n)" "ball 0 B2 \ cbox a (b::'n)" by auto obtain w where w: - "((\x. if x \ s then f x else 0) has_integral w) {a..b}" + "((\x. if x \ s then f x else 0) has_integral w) (cbox a b)" "norm (w - k) < e / 2" using B1(2)[OF *(1)] by blast obtain z where z: - "((\x. if x \ s then g x else 0) has_integral z) {a..b}" + "((\x. if x \ s then g x else 0) has_integral z) (cbox a b)" "norm (z - l) < e / 2" using B2(2)[OF *(2)] by blast have *: "\x. (if x \ s then f x + g x else 0) = (if x \ s then f x else 0) + (if x \ s then g x else 0)" by auto - show "\z. ((\x. if x \ s then f x + g x else 0) has_integral z) {a..b} \ norm (z - (k + l)) < e" + show "\z. ((\x. if x \ s then f x + g x else 0) has_integral z) (cbox a b) \ norm (z - (k + l)) < e" apply (rule_tac x="w + z" in exI) apply rule apply (rule lem[OF w(1) z(1), unfolded *[symmetric]]) @@ -2808,7 +2895,7 @@ by auto lemma integral_0: - "integral s (\x::'n::ordered_euclidean_space. 0::'m::real_normed_vector) = 0" + "integral s (\x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" by (rule integral_unique has_integral_0)+ lemma integral_add: "f integrable_on s \ g integrable_on s \ @@ -2880,7 +2967,7 @@ done lemma integral_component_eq[simp]: - fixes f :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on s" shows "integral s (\x. f x \ k) = integral s f \ k" unfolding integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] .. @@ -2938,8 +3025,8 @@ by auto lemma has_integral_null[dest]: - assumes "content({a..b}) = 0" - shows "(f has_integral 0) ({a..b})" + assumes "content(cbox a b) = 0" + shows "(f has_integral 0) (cbox a b)" unfolding has_integral apply rule apply rule @@ -2955,7 +3042,7 @@ then show "gauge (\x. ball x 1)" by auto fix p - assume p: "p tagged_division_of {a..b}" + assume p: "p tagged_division_of (cbox a b)" have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) = 0" unfolding norm_eq_zero diff_0_right using setsum_content_null[OF assms(1) p, of f] . @@ -2963,7 +3050,12 @@ using e by auto qed -lemma has_integral_null_eq[simp]: "content {a..b} = 0 \ (f has_integral i) {a..b} \ i = 0" +lemma has_integral_null_real[dest]: + assumes "content {a .. b::real} = 0" + shows "(f has_integral 0) {a .. b}" + by (metis assms box_real(2) has_integral_null) + +lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \ (f has_integral i) (cbox a b) \ i = 0" apply rule apply (rule has_integral_unique) apply assumption @@ -2972,13 +3064,13 @@ apply auto done -lemma integral_null[dest]: "content {a..b} = 0 \ integral {a..b} f = 0" +lemma integral_null[dest]: "content (cbox a b) = 0 \ integral (cbox a b) f = 0" apply (rule integral_unique) apply (drule has_integral_null) apply assumption done -lemma integrable_on_null[dest]: "content {a..b} = 0 \ f integrable_on {a..b}" +lemma integrable_on_null[dest]: "content (cbox a b) = 0 \ f integrable_on (cbox a b)" unfolding integrable_on_def apply (drule has_integral_null) apply auto @@ -3006,32 +3098,32 @@ by (rule integral_unique) (rule has_integral_empty) lemma has_integral_refl[intro]: - fixes a :: "'a::ordered_euclidean_space" - shows "(f has_integral 0) {a..a}" + fixes a :: "'a::euclidean_space" + shows "(f has_integral 0) (cbox a a)" and "(f has_integral 0) {a}" proof - - have *: "{a} = {a..a}" + have *: "{a} = cbox a a" apply (rule set_eqI) - unfolding mem_interval singleton_iff euclidean_eq_iff[where 'a='a] + unfolding mem_box singleton_iff euclidean_eq_iff[where 'a='a] apply safe prefer 3 apply (erule_tac x=b in ballE) apply (auto simp add: field_simps) done - show "(f has_integral 0) {a..a}" "(f has_integral 0) {a}" + show "(f has_integral 0) (cbox a a)" "(f has_integral 0) {a}" unfolding * apply (rule_tac[!] has_integral_null) unfolding content_eq_0_interior - unfolding interior_closed_interval - using interval_sing - apply auto - done -qed - -lemma integrable_on_refl[intro]: "f integrable_on {a..a}" + unfolding interior_cbox + using box_sing + apply auto + done +qed + +lemma integrable_on_refl[intro]: "f integrable_on cbox a a" unfolding integrable_on_def by auto -lemma integral_refl: "integral {a..a} f = 0" +lemma integral_refl: "integral (cbox a a) f = 0" by (rule integral_unique) auto @@ -3039,11 +3131,11 @@ (* XXXXXXX *) lemma integrable_cauchy: - fixes f :: "'n::ordered_euclidean_space \ 'a::{real_normed_vector,complete_space}" - shows "f integrable_on {a..b} \ + fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" + shows "f integrable_on cbox a b \ (\e>0.\d. gauge d \ - (\p1 p2. p1 tagged_division_of {a..b} \ d fine p1 \ - p2 tagged_division_of {a..b} \ d fine p2 \ + (\p1 p2. p1 tagged_division_of (cbox a b) \ d fine p1 \ + p2 tagged_division_of (cbox a b) \ d fine p2 \ norm (setsum (\(x,k). content k *\<^sub>R f x) p1 - setsum (\(x,k). content k *\<^sub>R f x) p2) < e))" (is "?l = (\e>0. \d. ?P e d)") @@ -3070,8 +3162,8 @@ apply (erule conjE)+ proof - fix p1 p2 - assume as: "p1 tagged_division_of {a..b}" "d fine p1" - "p2 tagged_division_of {a..b}" "d fine p2" + assume as: "p1 tagged_division_of (cbox a b)" "d fine p1" + "p2 tagged_division_of (cbox a b)" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" apply (rule dist_triangle_half_l[where y=y,unfolded dist_norm]) using d(2)[OF conjI[OF as(1-2)]] d(2)[OF conjI[OF as(3-4)]] . @@ -3087,7 +3179,7 @@ using d(1) apply auto done - then have "\n. \p. p tagged_division_of {a..b} \ (\x. \{d i x |i. i \ {0..n}}) fine p" + then have "\n. \p. p tagged_division_of (cbox a b) \ (\x. \{d i x |i. i \ {0..n}}) fine p" apply - proof case goal1 @@ -3133,7 +3225,7 @@ by auto guess N2 using y[OF *] .. note N2=this show "\d. gauge d \ - (\p. p tagged_division_of {a..b} \ d fine p \ + (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - y) < e)" apply (rule_tac x="d (N1 + N2)" in exI) apply rule @@ -3142,7 +3234,7 @@ show "gauge (d (N1 + N2))" using d by auto fix q - assume as: "q tagged_division_of {a..b}" "d (N1 + N2) fine q" + assume as: "q tagged_division_of (cbox a b)" "d (N1 + N2) fine q" have *: "inverse (real (N1 + N2 + 1)) < e / 2" apply (rule less_trans) using N1 @@ -3167,23 +3259,23 @@ subsection {* Additivity of integral on abutting intervals. *} lemma interval_split: - fixes a :: "'a::ordered_euclidean_space" + fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows - "{a..b} \ {x. x\k \ c} = {a .. (\i\Basis. (if i = k then min (b\k) c else b\i) *\<^sub>R i)}" - "{a..b} \ {x. x\k \ c} = {(\i\Basis. (if i = k then max (a\k) c else a\i) *\<^sub>R i) .. b}" + "cbox a b \ {x. x\k \ c} = cbox a (\i\Basis. (if i = k then min (b\k) c else b\i) *\<^sub>R i)" + "cbox a b \ {x. x\k \ c} = cbox (\i\Basis. (if i = k then max (a\k) c else a\i) *\<^sub>R i) b" apply (rule_tac[!] set_eqI) - unfolding Int_iff mem_interval mem_Collect_eq + unfolding Int_iff mem_box mem_Collect_eq using assms apply auto done lemma content_split: - fixes a :: "'a::ordered_euclidean_space" + fixes a :: "'a::euclidean_space" assumes "k \ Basis" - shows "content {a..b} = content({a..b} \ {x. x\k \ c}) + content({a..b} \ {x. x\k \ c})" + shows "content (cbox a b) = content(cbox a b \ {x. x\k \ c}) + content(cbox a b \ {x. x\k \ c})" proof cases - note simps = interval_split[OF assms] content_closed_interval_cases eucl_le[where 'a='a] + note simps = interval_split[OF assms] content_cbox_cases have *: "Basis = insert k (Basis - {k})" "\x. finite (Basis-{x})" "\x. x\Basis-{x}" using assms by auto have *: "\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" @@ -3194,7 +3286,7 @@ unfolding setprod_insert[OF *(2-)] apply auto done - assume as: "a \ b" + assume as: "\i\Basis. a \ i \ b \ i" moreover have "\x. min (b \ k) c = max (a \ k) c \ x * (b\k - a\k) = x * (max (a \ k) c - a \ k) + x * (b \ k - max (a \ k) c)" @@ -3207,7 +3299,7 @@ by (auto intro!: setprod_cong) have "\ a \ k \ c \ \ c \ b \ k \ False" unfolding not_le - using as[unfolded eucl_le[where 'a='a],rule_format,of k] assms + using as[unfolded ,rule_format,of k] assms by auto ultimately show ?thesis using assms @@ -3216,15 +3308,15 @@ unfolding *(2) by auto next - assume "\ a \ b" - then have "{a .. b} = {}" - unfolding interval_eq_empty by (auto simp: eucl_le[where 'a='a] not_le) + assume "\ (\i\Basis. a \ i \ b \ i)" + then have "cbox a b = {}" + unfolding box_eq_empty by (auto simp: not_le) then show ?thesis by (auto simp: not_le) qed lemma division_split_left_inj: - fixes type :: "'a::ordered_euclidean_space" + fixes type :: "'a::euclidean_space" assumes "d division_of i" and "k1 \ d" and "k2 \ d" @@ -3234,8 +3326,8 @@ shows "content(k1 \ {x. x\k \ c}) = 0" proof - note d=division_ofD[OF assms(1)] - have *: "\(a::'a) b c. content ({a..b} \ {x. x\k \ c}) = 0 \ - interior({a..b} \ {x. x\k \ c}) = {}" + have *: "\(a::'a) b c. content (cbox a b \ {x. x\k \ c}) = 0 \ + interior(cbox a b \ {x. x\k \ c}) = {}" unfolding interval_split[OF k] content_eq_0_interior by auto guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this @@ -3252,7 +3344,7 @@ qed lemma division_split_right_inj: - fixes type :: "'a::ordered_euclidean_space" + fixes type :: "'a::euclidean_space" assumes "d division_of i" and "k1 \ d" and "k2 \ d" @@ -3262,8 +3354,8 @@ shows "content (k1 \ {x. x\k \ c}) = 0" proof - note d=division_ofD[OF assms(1)] - have *: "\a b::'a. \c. content({a..b} \ {x. x\k \ c}) = 0 \ - interior({a..b} \ {x. x\k \ c}) = {}" + have *: "\a b::'a. \c. content(cbox a b \ {x. x\k \ c}) = 0 \ + interior(cbox a b \ {x. x\k \ c}) = {}" unfolding interval_split[OF k] content_eq_0_interior by auto guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this @@ -3280,7 +3372,7 @@ qed lemma tagged_division_split_left_inj: - fixes x1 :: "'a::ordered_euclidean_space" + fixes x1 :: "'a::euclidean_space" assumes "d tagged_division_of i" and "(x1, k1) \ d" and "(x2, k2) \ d" @@ -3303,7 +3395,7 @@ qed lemma tagged_division_split_right_inj: - fixes x1 :: "'a::ordered_euclidean_space" + fixes x1 :: "'a::euclidean_space" assumes "d tagged_division_of i" and "(x1, k1) \ d" and "(x2, k2) \ d" @@ -3326,12 +3418,12 @@ qed lemma division_split: - fixes a :: "'a::ordered_euclidean_space" - assumes "p division_of {a..b}" + fixes a :: "'a::euclidean_space" + assumes "p division_of (cbox a b)" and k: "k\Basis" - shows "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of({a..b} \ {x. x\k \ c})" + shows "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of(cbox a b \ {x. x\k \ c})" (is "?p1 division_of ?I1") - and "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of ({a..b} \ {x. x\k \ c})" + and "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of (cbox a b \ {x. x\k \ c})" (is "?p2 division_of ?I2") proof (rule_tac[!] division_ofI) note p = division_ofD[OF assms(1)] @@ -3344,7 +3436,7 @@ assume "k \ ?p1" then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this guess u v using p(4)[OF l(2)] by (elim exE) note uv=this - show "k \ ?I1" "k \ {}" "\a b. k = {a..b}" + show "k \ ?I1" "k \ {}" "\a b. k = cbox a b" unfolding l using p(2-3)[OF l(2)] l(3) unfolding uv @@ -3365,7 +3457,7 @@ assume "k \ ?p2" then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this guess u v using p(4)[OF l(2)] by (elim exE) note uv=this - show "k \ ?I2" "k \ {}" "\a b. k = {a..b}" + show "k \ ?I2" "k \ {}" "\a b. k = cbox a b" unfolding l using p(2-3)[OF l(2)] l(3) unfolding uv @@ -3384,11 +3476,11 @@ qed lemma has_integral_split: - fixes f :: "'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "(f has_integral i) ({a..b} \ {x. x\k \ c})" - and "(f has_integral j) ({a..b} \ {x. x\k \ c})" + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes "(f has_integral i) (cbox a b \ {x. x\k \ c})" + and "(f has_integral j) (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" - shows "(f has_integral (i + j)) ({a..b})" + shows "(f has_integral (i + j)) (cbox a b)" proof (unfold has_integral, rule, rule) case goal1 then have e: "e/2 > 0" @@ -3409,7 +3501,7 @@ show "gauge ?d" using d1(1) d2(1) unfolding gauge_def by auto fix p - assume "p tagged_division_of {a..b}" "?d fine p" + assume "p tagged_division_of (cbox a b)" "?d fine p" note p = this tagged_division_ofD[OF this(1)] have lem0: "\x kk. (x, kk) \ p \ kk \ {x. x\k \ c} \ {} \ x\k \ c" @@ -3499,7 +3591,7 @@ prefer 6 apply (rule fineI) proof - - show "\{k. \x. (x, k) \ ?M1} = {a..b} \ {x. x\k \ c}" + show "\{k. \x. (x, k) \ ?M1} = cbox a b \ {x. x\k \ c}" unfolding p(8)[symmetric] by auto fix x l assume xl: "(x, l) \ ?M1" @@ -3510,11 +3602,11 @@ done then show "l \ d1 x" unfolding xl' by auto - show "x \ l" "l \ {a..b} \ {x. x \ k \ c}" + show "x \ l" "l \ cbox a b \ {x. x \ k \ c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4) using lem0(1)[OF xl'(3-4)] by auto - show "\a b. l = {a..b}" + show "\a b. l = cbox a b" unfolding xl' using p(6)[OF xl'(3)] by (fastforce simp add: interval_split[OF k,where c=c]) @@ -3544,7 +3636,7 @@ prefer 6 apply (rule fineI) proof - - show "\{k. \x. (x, k) \ ?M2} = {a..b} \ {x. x\k \ c}" + show "\{k. \x. (x, k) \ ?M2} = cbox a b \ {x. x\k \ c}" unfolding p(8)[symmetric] by auto fix x l assume xl: "(x, l) \ ?M2" @@ -3555,13 +3647,13 @@ done then show "l \ d2 x" unfolding xl' by auto - show "x \ l" "l \ {a..b} \ {x. x \ k \ c}" + show "x \ l" "l \ cbox a b \ {x. x \ k \ c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4) using lem0(2)[OF xl'(3-4)] by auto - show "\a b. l = {a..b}" + show "\a b. l = cbox a b" unfolding xl' using p(6)[OF xl'(3)] by (fastforce simp add: interval_split[OF k, where c=c]) @@ -3654,31 +3746,41 @@ subsection {* A sort of converse, integrability on subintervals. *} lemma tagged_division_union_interval: - fixes a :: "'a::ordered_euclidean_space" - assumes "p1 tagged_division_of ({a..b} \ {x. x\k \ (c::real)})" - and "p2 tagged_division_of ({a..b} \ {x. x\k \ c})" + fixes a :: "'a::euclidean_space" + assumes "p1 tagged_division_of (cbox a b \ {x. x\k \ (c::real)})" + and "p2 tagged_division_of (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" - shows "(p1 \ p2) tagged_division_of ({a..b})" -proof - - have *: "{a..b} = ({a..b} \ {x. x\k \ c}) \ ({a..b} \ {x. x\k \ c})" + shows "(p1 \ p2) tagged_division_of (cbox a b)" +proof - + have *: "cbox a b = (cbox a b \ {x. x\k \ c}) \ (cbox a b \ {x. x\k \ c})" by auto show ?thesis apply (subst *) apply (rule tagged_division_union[OF assms(1-2)]) - unfolding interval_split[OF k] interior_closed_interval + unfolding interval_split[OF k] interior_cbox using k - apply (auto simp add: interval elim!: ballE[where x=k]) - done -qed + apply (auto simp add: box elim!: ballE[where x=k]) + done +qed + +lemma tagged_division_union_interval_real: + fixes a :: real + assumes "p1 tagged_division_of ({a .. b} \ {x. x\k \ (c::real)})" + and "p2 tagged_division_of ({a .. b} \ {x. x\k \ c})" + and k: "k \ Basis" + shows "(p1 \ p2) tagged_division_of {a .. b}" + using assms + unfolding box_real[symmetric] + by (rule tagged_division_union_interval) lemma has_integral_separate_sides: - fixes f :: "'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "(f has_integral i) ({a..b})" + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes "(f has_integral i) (cbox a b)" and "e > 0" and k: "k \ Basis" obtains d where "gauge d" - "\p1 p2. p1 tagged_division_of ({a..b} \ {x. x\k \ c}) \ d fine p1 \ - p2 tagged_division_of ({a..b} \ {x. x\k \ c}) \ d fine p2 \ + "\p1 p2. p1 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p1 \ + p2 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p2 \ norm ((setsum (\(x,k). content k *\<^sub>R f x) p1 + setsum (\(x,k). content k *\<^sub>R f x) p2) - i) < e" proof - guess d using has_integralD[OF assms(1-2)] . note d=this @@ -3691,9 +3793,9 @@ apply (elim conjE) proof - fix p1 p2 - assume "p1 tagged_division_of {a..b} \ {x. x \ k \ c}" "d fine p1" + assume "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this - assume "p2 tagged_division_of {a..b} \ {x. c \ x \ k}" "d fine p2" + assume "p2 tagged_division_of (cbox a b) \ {x. c \ x \ k}" "d fine p2" note p2=tagged_division_ofD[OF this(1)] this note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = @@ -3755,11 +3857,11 @@ qed lemma integrable_split[intro]: - fixes f :: "'a::ordered_euclidean_space \ 'b::{real_normed_vector,complete_space}" - assumes "f integrable_on {a..b}" + fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" + assumes "f integrable_on cbox a b" and k: "k \ Basis" - shows "f integrable_on ({a..b} \ {x. x\k \ c})" (is ?t1) - and "f integrable_on ({a..b} \ {x. x\k \ c})" (is ?t2) + shows "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?t1) + and "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?t2) proof - guess y using assms(1) unfolding integrable_on_def .. note y=this def b' \ "\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i::'a" @@ -3773,8 +3875,8 @@ then have "e/2>0" by auto from has_integral_separate_sides[OF y this k,of c] guess d . note d=this[rule_format] - let ?P = "\A. \d. gauge d \ (\p1 p2. p1 tagged_division_of {a..b} \ A \ d fine p1 \ - p2 tagged_division_of {a..b} \ A \ d fine p2 \ + let ?P = "\A. \d. gauge d \ (\p1 p2. p1 tagged_division_of (cbox a b) \ A \ d fine p1 \ + p2 tagged_division_of (cbox a b) \ A \ d fine p2 \ norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e)" show "?P {x. x \ k \ c}" apply (rule_tac x=d in exI) @@ -3785,8 +3887,8 @@ apply rule proof - fix p1 p2 - assume as: "p1 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p1 \ - p2 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p2" + assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c} \ d fine p1 \ + p2 tagged_division_of (cbox a b) \ {x. x \ k \ c} \ d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof - guess p using fine_division_exists[OF d(1), of a' b] . note p=this @@ -3806,8 +3908,8 @@ apply rule proof - fix p1 p2 - assume as: "p1 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p1 \ - p2 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p2" + assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c} \ d fine p1 \ + p2 tagged_division_of (cbox a b) \ {x. x \ k \ c} \ d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof - guess p using fine_division_exists[OF d(1), of a b'] . note p=this @@ -3828,23 +3930,23 @@ definition "neutral opp = (SOME x. \y. opp x y = y \ opp y x = y)" -definition operative :: "('a \ 'a \ 'a) \ (('b::ordered_euclidean_space) set \ 'a) \ bool" +definition operative :: "('a \ 'a \ 'a) \ (('b::euclidean_space) set \ 'a) \ bool" where "operative opp f \ - (\a b. content {a..b} = 0 \ f {a..b} = neutral opp) \ - (\a b c. \k\Basis. f {a..b} = opp (f({a..b} \ {x. x\k \ c})) (f ({a..b} \ {x. x\k \ c})))" + (\a b. content (cbox a b) = 0 \ f (cbox a b) = neutral opp) \ + (\a b c. \k\Basis. f (cbox a b) = opp (f(cbox a b \ {x. x\k \ c})) (f (cbox a b \ {x. x\k \ c})))" lemma operativeD[dest]: - fixes type :: "'a::ordered_euclidean_space" + fixes type :: "'a::euclidean_space" assumes "operative opp f" - shows "\a b::'a. content {a..b} = 0 \ f {a..b} = neutral opp" - and "\a b c k. k \ Basis \ f {a..b} = - opp (f ({a..b} \ {x. x\k \ c})) (f ({a..b} \ {x. x\k \ c}))" + shows "\a b::'a. content (cbox a b) = 0 \ f (cbox a b) = neutral opp" + and "\a b c k. k \ Basis \ f (cbox a b) = + opp (f (cbox a b \ {x. x\k \ c})) (f (cbox a b \ {x. x\k \ c}))" using assms unfolding operative_def by auto -lemma operative_trivial: "operative opp f \ content {a..b} = 0 \ f {a..b} = neutral opp" +lemma operative_trivial: "operative opp f \ content (cbox a b) = 0 \ f (cbox a b) = neutral opp" unfolding operative_def by auto -lemma property_empty_interval: "\a b. content {a..b} = 0 \ P {a..b} \ P {}" +lemma property_empty_interval: "\a b. content (cbox a b) = 0 \ P (cbox a b) \ P {}" using content_empty unfolding empty_as_interval by auto lemma operative_empty: "operative opp f \ f {} = neutral opp" @@ -4042,7 +4144,7 @@ by (auto simp add: algebra_simps) lemma operative_integral: - fixes f :: "'a::ordered_euclidean_space \ 'b::banach" + fixes f :: "'a::euclidean_space \ 'b::banach" shows "operative (lifted(op +)) (\i. if f integrable_on i then Some(integral i f) else None)" unfolding operative_def unfolding neutral_lifted[OF monoidal_monoid] neutral_add @@ -4056,10 +4158,10 @@ fix a b c fix k :: 'a assume k: "k \ Basis" - show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) = - lifted op + (if f integrable_on {a..b} \ {x. x \ k \ c} then Some (integral ({a..b} \ {x. x \ k \ c}) f) else None) - (if f integrable_on {a..b} \ {x. c \ x \ k} then Some (integral ({a..b} \ {x. c \ x \ k}) f) else None)" - proof (cases "f integrable_on {a..b}") + show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = + lifted op + (if f integrable_on cbox a b \ {x. x \ k \ c} then Some (integral (cbox a b \ {x. x \ k \ c}) f) else None) + (if f integrable_on cbox a b \ {x. c \ x \ k} then Some (integral (cbox a b \ {x. c \ x \ k}) f) else None)" + proof (cases "f integrable_on cbox a b") case True show ?thesis unfolding if_P[OF True] @@ -4076,13 +4178,13 @@ done next case False - have "\ (f integrable_on {a..b} \ {x. x \ k \ c}) \ \ ( f integrable_on {a..b} \ {x. c \ x \ k})" + have "\ (f integrable_on cbox a b \ {x. x \ k \ c}) \ \ ( f integrable_on cbox a b \ {x. c \ x \ k})" proof (rule ccontr) assume "\ ?thesis" - then have "f integrable_on {a..b}" + then have "f integrable_on cbox a b" apply - unfolding integrable_on_def - apply (rule_tac x="integral ({a..b} \ {x. x \ k \ c}) f + integral ({a..b} \ {x. x \ k \ c}) f" in exI) + apply (rule_tac x="integral (cbox a b \ {x. x \ k \ c}) f + integral (cbox a b \ {x. x \ k \ c}) f" in exI) apply (rule has_integral_split[OF _ _ k]) apply (rule_tac[!] integrable_integral) apply auto @@ -4095,8 +4197,8 @@ qed next fix a b :: 'a - assume as: "content {a..b} = 0" - then show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) = Some 0" + assume as: "content (cbox a b) = 0" + then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0" unfolding if_P[OF integrable_on_null[OF as]] using has_integral_null_eq[OF as] by auto @@ -4105,18 +4207,18 @@ subsection {* Points of division of a partition. *} -definition "division_points (k::('a::ordered_euclidean_space) set) d = - {(j,x). j \ Basis \ (Inf k)\j < x \ x < (Sup k)\j \ - (\i\d. (Inf i)\j = x \ (Sup i)\j = x)}" +definition "division_points (k::('a::euclidean_space) set) d = + {(j,x). j \ Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ + (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" lemma division_points_finite: - fixes i :: "'a::ordered_euclidean_space set" + fixes i :: "'a::euclidean_space set" assumes "d division_of i" shows "finite (division_points i d)" proof - note assm = division_ofD[OF assms] - let ?M = "\j. {(j,x)|x. (Inf i)\j < x \ x < (Sup i)\j \ - (\i\d. (Inf i)\j = x \ (Sup i)\j = x)}" + let ?M = "\j. {(j,x)|x. (interval_lowerbound i)\j < x \ x < (interval_upperbound i)\j \ + (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" have *: "division_points i d = \(?M ` Basis)" unfolding division_points_def by auto show ?thesis @@ -4124,14 +4226,14 @@ qed lemma division_points_subset: - fixes a :: "'a::ordered_euclidean_space" - assumes "d division_of {a..b}" + fixes a :: "'a::euclidean_space" + assumes "d division_of (cbox a b)" and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" and k: "k \ Basis" - shows "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ l \ {x. x\k \ c} \ {}} \ - division_points ({a..b}) d" (is ?t1) - and "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} \ - division_points ({a..b}) d" (is ?t2) + shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ l \ {x. x\k \ c} \ {}} \ + division_points (cbox a b) d" (is ?t1) + and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} \ + division_points (cbox a b) d" (is ?t2) proof - note assm = division_ofD[OF assms(1)] have *: "\i\Basis. a\i \ b\i" @@ -4153,15 +4255,15 @@ fix i l x assume as: "a \ fst x < snd x" "snd x < (if fst x = k then c else b \ fst x)" - "Inf i \ fst x = snd x \ Sup i \ fst x = snd x" + "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" "i = l \ {x. x \ k \ c}" "l \ d" "l \ {x. x \ k \ c} \ {}" and fstx: "fst x \ Basis" from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this have *: "\i\Basis. u \ i \ (\i\Basis. (if i = k then min (v \ k) c else v \ i) *\<^sub>R i) \ i" - using as(6) unfolding l interval_split[OF k] interval_ne_empty as . + using as(6) unfolding l interval_split[OF k] box_ne_empty as . have **: "\i\Basis. u\i \ v\i" - using l using as(6) unfolding interval_ne_empty[symmetric] by auto - show "\i\d. Inf i \ fst x = snd x \ Sup i \ fst x = snd x" + using l using as(6) unfolding box_ne_empty[symmetric] by auto + show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" apply (rule bexI[OF _ `l \ d`]) using as(1-3,5) fstx unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as @@ -4184,15 +4286,15 @@ fix i l x assume as: "(if fst x = k then c else a \ fst x) < snd x" "snd x < b \ fst x" - "Inf i \ fst x = snd x \ Sup i \ fst x = snd x" + "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" "i = l \ {x. c \ x \ k}" "l \ d" "l \ {x. c \ x \ k} \ {}" and fstx: "fst x \ Basis" from assm(4)[OF this(5)] guess u v by (elim exE) note l=this have *: "\i\Basis. (\i\Basis. (if i = k then max (u \ k) c else u \ i) *\<^sub>R i) \ i \ v \ i" - using as(6) unfolding l interval_split[OF k] interval_ne_empty as . + using as(6) unfolding l interval_split[OF k] box_ne_empty as . have **: "\i\Basis. u\i \ v\i" - using l using as(6) unfolding interval_ne_empty[symmetric] by auto - show "\i\d. Inf i \ fst x = snd x \ Sup i \ fst x = snd x" + using l using as(6) unfolding box_ne_empty[symmetric] by auto + show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" apply (rule bexI[OF _ `l \ d`]) using as(1-3,5) fstx unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as @@ -4204,32 +4306,32 @@ qed lemma division_points_psubset: - fixes a :: "'a::ordered_euclidean_space" - assumes "d division_of {a..b}" + fixes a :: "'a::euclidean_space" + assumes "d division_of (cbox a b)" and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" and "l \ d" - and "Inf l\k = c \ Sup l\k = c" + and "interval_lowerbound l\k = c \ interval_upperbound l\k = c" and k: "k \ Basis" - shows "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ - division_points ({a..b}) d" (is "?D1 \ ?D") - and "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ - division_points ({a..b}) d" (is "?D2 \ ?D") + shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ + division_points (cbox a b) d" (is "?D1 \ ?D") + and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ + division_points (cbox a b) d" (is "?D2 \ ?D") proof - have ab: "\i\Basis. a\i \ b\i" using assms(2) by (auto intro!:less_imp_le) guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this have uv: "\i\Basis. u\i \ v\i" "\i\Basis. a\i \ u\i \ v\i \ b\i" - using division_ofD(2,2,3)[OF assms(1,5)] unfolding l interval_ne_empty + using division_ofD(2,2,3)[OF assms(1,5)] unfolding l box_ne_empty unfolding subset_eq apply - defer apply (erule_tac x=u in ballE) apply (erule_tac x=v in ballE) - unfolding mem_interval - apply auto - done - have *: "Sup ({a..b} \ {x. x \ k \ Sup l \ k}) \ k = Sup l \ k" - "Sup ({a..b} \ {x. x \ k \ Inf l \ k}) \ k = Inf l \ k" + unfolding mem_box + apply auto + done + have *: "interval_upperbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" + "interval_upperbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" unfolding interval_split[OF k] apply (subst interval_bounds) prefer 3 @@ -4242,9 +4344,9 @@ using assms(2-) apply - apply (erule disjE) - apply (rule_tac x="(k,(Inf l)\k)" in exI) + apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI) defer - apply (rule_tac x="(k,(Sup l)\k)" in exI) + apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI) unfolding division_points_def unfolding interval_bounds[OF ab] apply (auto simp add:*) @@ -4257,8 +4359,8 @@ apply auto done - have *: "Inf ({a..b} \ {x. x \ k \ Inf l \ k}) \ k = Inf l \ k" - "Inf ({a..b} \ {x. x \ k \ Sup l \ k}) \ k = Sup l \ k" + have *: "interval_lowerbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" + "interval_lowerbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" unfolding interval_split[OF k] apply (subst interval_bounds) prefer 3 @@ -4271,9 +4373,9 @@ using assms(2-) apply - apply (erule disjE) - apply (rule_tac x="(k,(Inf l)\k)" in exI) + apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI) defer - apply (rule_tac x="(k,(Sup l)\k)" in exI) + apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI) unfolding division_points_def unfolding interval_bounds[OF ab] apply (auto simp add:*) @@ -4459,25 +4561,25 @@ using assms by auto lemma operative_division: - fixes f :: "'a::ordered_euclidean_space set \ 'b" + fixes f :: "'a::euclidean_space set \ 'b" assumes "monoidal opp" and "operative opp f" - and "d division_of {a..b}" - shows "iterate opp d f = f {a..b}" -proof - - def C \ "card (division_points {a..b} d)" + and "d division_of (cbox a b)" + shows "iterate opp d f = f (cbox a b)" +proof - + def C \ "card (division_points (cbox a b) d)" then show ?thesis using assms proof (induct C arbitrary: a b d rule: full_nat_induct) case goal1 - { presume *: "content {a..b} \ 0 \ ?case" + { presume *: "content (cbox a b) \ 0 \ ?case" then show ?case apply - apply cases defer apply assumption proof - - assume as: "content {a..b} = 0" + assume as: "content (cbox a b) = 0" show ?case unfolding operativeD(1)[OF assms(2) as] apply(rule iterate_eq_neutral[OF goal1(2)]) @@ -4495,13 +4597,13 @@ qed qed } - assume "content {a..b} \ 0" note ab = this[unfolded content_lt_nz[symmetric] content_pos_lt_eq] + assume "content (cbox a b) \ 0" note ab = this[unfolded content_lt_nz[symmetric] content_pos_lt_eq] then have ab': "\i\Basis. a\i \ b\i" by (auto intro!: less_imp_le) show ?case - proof (cases "division_points {a..b} d = {}") + proof (cases "division_points (cbox a b) d = {}") case True - have d': "\i\d. \u v. i = {u..v} \ + have d': "\i\d. \u v. i = cbox u v \ (\j\Basis. u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j)" unfolding forall_in_division[OF goal1(4)] apply rule @@ -4515,14 +4617,14 @@ fix u v fix j :: 'a assume j: "j \ Basis" - assume as: "{u..v} \ d" + assume as: "cbox u v \ d" note division_ofD(3)[OF goal1(4) this] then have uv: "\i\Basis. u\i \ v\i" "u\j \ v\j" - using j unfolding interval_ne_empty by auto - have *: "\p r Q. \ j\Basis \ p \ r \ (\x\d. Q x) \ p \ r \ Q {u..v}" + using j unfolding box_ne_empty by auto + have *: "\p r Q. \ j\Basis \ p \ r \ (\x\d. Q x) \ p \ r \ Q (cbox u v)" using as j by auto - have "(j, u\j) \ division_points {a..b} d" - "(j, v\j) \ division_points {a..b} d" using True by auto + have "(j, u\j) \ division_points (cbox a b) d" + "(j, v\j) \ division_points (cbox a b) d" using True by auto note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps] note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]] moreover @@ -4531,39 +4633,39 @@ unfolding subset_eq apply - apply (erule_tac x=u in ballE,erule_tac[3] x=v in ballE) - unfolding interval_ne_empty mem_interval + unfolding box_ne_empty mem_box using j apply auto done ultimately show "u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j" unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by auto qed - have "(1/2) *\<^sub>R (a+b) \ {a..b}" - unfolding mem_interval using ab by(auto intro!: less_imp_le simp: inner_simps) + have "(1/2) *\<^sub>R (a+b) \ cbox a b" + unfolding mem_box using ab by(auto intro!: less_imp_le simp: inner_simps) note this[unfolded division_ofD(6)[OF goal1(4),symmetric] Union_iff] then guess i .. note i=this guess u v using d'[rule_format,OF i(1)] by (elim exE conjE) note uv=this - have "{a..b} \ d" + have "cbox a b \ d" proof - - { presume "i = {a..b}" then show ?thesis using i by auto } - { presume "u = a" "v = b" then show "i = {a..b}" using uv by auto } + { presume "i = cbox a b" then show ?thesis using i by auto } + { presume "u = a" "v = b" then show "i = cbox a b" using uv by auto } show "u = a" "v = b" unfolding euclidean_eq_iff[where 'a='a] proof safe fix j :: 'a assume j: "j \ Basis" - note i(2)[unfolded uv mem_interval,rule_format,of j] + note i(2)[unfolded uv mem_box,rule_format,of j] then show "u \ j = a \ j" and "v \ j = b \ j" using uv(2)[rule_format,of j] j by (auto simp: inner_simps) qed qed - then have *: "d = insert {a..b} (d - {{a..b}})" + then have *: "d = insert (cbox a b) (d - {cbox a b})" by auto - have "iterate opp (d - {{a..b}}) f = neutral opp" + have "iterate opp (d - {cbox a b}) f = neutral opp" apply (rule iterate_eq_neutral[OF goal1(2)]) proof fix x - assume x: "x \ d - {{a..b}}" + assume x: "x \ d - {cbox a b}" then have "x\d" by auto note d'[rule_format,OF this] then guess u v by (elim exE conjE) note uv=this @@ -4573,7 +4675,7 @@ unfolding euclidean_eq_iff[where 'a='a] by auto then have "u\j = v\j" using uv(2)[rule_format,OF j] by auto - then have "content {u..v} = 0" + then have "content (cbox u v) = 0" unfolding content_eq_0 apply (rule_tac x=j in bexI) using j @@ -4582,7 +4684,7 @@ then show "f x = neutral opp" unfolding uv(1) by (rule operativeD(1)[OF goal1(3)]) qed - then show "iterate opp d f = f {a..b}" + then show "iterate opp d f = f (cbox a b)" apply - apply (subst *) apply (subst iterate_insert[OF goal1(2)]) @@ -4591,7 +4693,7 @@ done next case False - then have "\x. x \ division_points {a..b} d" + then have "\x. x \ division_points (cbox a b) d" by auto then guess k c unfolding split_paired_Ex @@ -4606,8 +4708,8 @@ def ca \ "(\i\Basis. (if i = k then c else a\i) *\<^sub>R i)::'a" note division_points_psubset[OF goal1(4) ab kc(1-2) j] note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)] - then have *: "(iterate opp d1 f) = f ({a..b} \ {x. x\k \ c})" - "(iterate opp d2 f) = f ({a..b} \ {x. x\k \ c})" + then have *: "(iterate opp d1 f) = f (cbox a b \ {x. x\k \ c})" + "(iterate opp d2 f) = f (cbox a b \ {x. x\k \ c})" unfolding interval_split[OF kc(4)] apply (rule_tac[!] goal1(1)[rule_format]) using division_split[OF goal1(4), where k=k and c=c] @@ -4618,7 +4720,7 @@ using kc(4) apply auto done - have "f {a..b} = opp (iterate opp d1 f) (iterate opp d2 f)" (is "_ = ?prev") + have "f (cbox a b) = opp (iterate opp d1 f) (iterate opp d2 f)" (is "_ = ?prev") unfolding * apply (rule operativeD(2)) using goal1(3) @@ -4728,8 +4830,8 @@ lemma operative_tagged_division: assumes "monoidal opp" and "operative opp f" - and "d tagged_division_of {a..b}" - shows "iterate opp d (\(x,l). f l) = f {a..b}" + and "d tagged_division_of (cbox a b)" + shows "iterate opp d (\(x,l). f l) = f (cbox a b)" proof - have *: "(\(x,l). f l) = f \ snd" unfolding o_def by rule auto note assm = tagged_division_ofD[OF assms(3)] @@ -4752,7 +4854,7 @@ apply auto done qed - also have "\ = f {a..b}" + also have "\ = f (cbox a b)" using operative_division[OF assms(1-2) division_of_tagged_division[OF assms(3)]] . finally show ?thesis . qed @@ -4775,8 +4877,8 @@ qed lemma additive_content_division: - assumes "d division_of {a..b}" - shows "setsum content d = content {a..b}" + assumes "d division_of (cbox a b)" + shows "setsum content d = content (cbox a b)" unfolding operative_division[OF monoidal_monoid operative_content assms,symmetric] apply (subst setsum_iterate) using assms @@ -4784,8 +4886,8 @@ done lemma additive_content_tagged_division: - assumes "d tagged_division_of {a..b}" - shows "setsum (\(x,l). content l) d = content {a..b}" + assumes "d tagged_division_of (cbox a b)" + shows "setsum (\(x,l). content l) d = content (cbox a b)" unfolding operative_tagged_division[OF monoidal_monoid operative_content assms,symmetric] apply (subst setsum_iterate) using assms @@ -4796,8 +4898,8 @@ subsection {* Finally, the integral of a constant *} lemma has_integral_const[intro]: - fixes a b :: "'a::ordered_euclidean_space" - shows "((\x. c) has_integral (content {a..b} *\<^sub>R c)) {a..b}" + fixes a b :: "'a::euclidean_space" + shows "((\x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" unfolding has_integral apply rule apply rule @@ -4815,18 +4917,28 @@ apply auto done +lemma has_integral_const_real[intro]: + fixes a b :: real + shows "((\x. c) has_integral (content {a .. b} *\<^sub>R c)) {a .. b}" + by (metis box_real(2) has_integral_const) + lemma integral_const[simp]: - fixes a b :: "'a::ordered_euclidean_space" + fixes a b :: "'a::euclidean_space" + shows "integral (cbox a b) (\x. c) = content (cbox a b) *\<^sub>R c" + by (rule integral_unique) (rule has_integral_const) + +lemma integral_const_real[simp]: + fixes a b :: real shows "integral {a .. b} (\x. c) = content {a .. b} *\<^sub>R c" - by (rule integral_unique) (rule has_integral_const) + by (metis box_real(2) integral_const) subsection {* Bounds on the norm of Riemann sums and the integral itself. *} lemma dsum_bound: - assumes "p division_of {a..b}" + assumes "p division_of (cbox a b)" and "norm c \ e" - shows "norm (setsum (\l. content l *\<^sub>R c) p) \ e * content({a..b})" + shows "norm (setsum (\l. content l *\<^sub>R c) p) \ e * content(cbox a b)" apply (rule order_trans) apply (rule norm_setsum) unfolding norm_scaleR setsum_left_distrib[symmetric] @@ -4850,10 +4962,10 @@ qed (insert assms, auto) lemma rsum_bound: - assumes "p tagged_division_of {a..b}" - and "\x\{a..b}. norm (f x) \ e" - shows "norm (setsum (\(x,k). content k *\<^sub>R f x) p) \ e * content {a..b}" -proof (cases "{a..b} = {}") + assumes "p tagged_division_of (cbox a b)" + and "\x\cbox a b. norm (f x) \ e" + shows "norm (setsum (\(x,k). content k *\<^sub>R f x) p) \ e * content (cbox a b)" +proof (cases "cbox a b = {}") case True show ?thesis using assms(1) unfolding True tagged_division_of_trivial by auto @@ -4875,7 +4987,7 @@ apply (subst o_def) apply (rule abs_of_nonneg) proof - - show "setsum (content \ snd) p \ content {a..b}" + show "setsum (content \ snd) p \ content (cbox a b)" apply (rule eq_refl) unfolding additive_content_tagged_division[OF assms(1),symmetric] split_def apply auto @@ -4901,10 +5013,10 @@ qed lemma rsum_diff_bound: - assumes "p tagged_division_of {a..b}" - and "\x\{a..b}. norm (f x - g x) \ e" + assumes "p tagged_division_of (cbox a b)" + and "\x\cbox a b. norm (f x - g x) \ e" shows "norm (setsum (\(x,k). content k *\<^sub>R f x) p - setsum (\(x,k). content k *\<^sub>R g x) p) \ - e * content {a..b}" + e * content (cbox a b)" apply (rule order_trans[OF _ rsum_bound[OF assms]]) apply (rule eq_refl) apply (rule arg_cong[where f=norm]) @@ -4915,19 +5027,19 @@ done lemma has_integral_bound: - fixes f :: "'a::ordered_euclidean_space \ 'b::real_normed_vector" + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" - and "(f has_integral i) {a..b}" - and "\x\{a..b}. norm (f x) \ B" - shows "norm i \ B * content {a..b}" -proof - - let ?P = "content {a..b} > 0" + and "(f has_integral i) (cbox a b)" + and "\x\cbox a b. norm (f x) \ B" + shows "norm i \ B * content (cbox a b)" +proof - + let ?P = "content (cbox a b) > 0" { presume "?P \ ?thesis" then show ?thesis proof (cases ?P) case False - then have *: "content {a..b} = 0" + then have *: "content (cbox a b) = 0" using content_lt_nz by auto then have **: "i = 0" using assms(2) @@ -4941,7 +5053,7 @@ assume ab: ?P { presume "\ ?thesis \ False" then show ?thesis by auto } assume "\ ?thesis" - then have *: "norm i - B * content {a..b} > 0" + then have *: "norm i - B * content (cbox a b) > 0" by auto from assms(2)[unfolded has_integral,rule_format,OF *] guess d by (elim exE conjE) note d=this[rule_format] @@ -4959,13 +5071,21 @@ using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto qed +lemma has_integral_bound_real: + fixes f :: "real \ 'b::real_normed_vector" + assumes "0 \ B" + and "(f has_integral i) {a .. b}" + and "\x\{a .. b}. norm (f x) \ B" + shows "norm i \ B * content {a .. b}" + by (metis assms(1) assms(2) assms(3) box_real(2) has_integral_bound) + subsection {* Similar theorems about relationship among components. *} lemma rsum_component_le: - fixes f :: "'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "p tagged_division_of {a..b}" - and "\x\{a..b}. (f x)\i \ (g x)\i" + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "p tagged_division_of (cbox a b)" + and "\x\cbox a b. (f x)\i \ (g x)\i" shows "(setsum (\(x,k). content k *\<^sub>R f x) p)\i \ (setsum (\(x,k). content k *\<^sub>R g x) p)\i" unfolding inner_setsum_left apply (rule setsum_mono) @@ -4987,14 +5107,14 @@ qed lemma has_integral_component_le: - fixes f g :: "'a::ordered_euclidean_space \ 'b::euclidean_space" + fixes f g :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" assumes "(f has_integral i) s" "(g has_integral j) s" and "\x\s. (f x)\k \ (g x)\k" shows "i\k \ j\k" proof - - have lem: "\a b i j::'b. \g f::'a \ 'b. (f has_integral i) {a..b} \ - (g has_integral j) ({a..b}) \ \x\{a..b}. (f x)\k \ (g x)\k \ i\k \ j\k" + have lem: "\a b i j::'b. \g f::'a \ 'b. (f has_integral i) (cbox a b) \ + (g has_integral j) (cbox a b) \ \x\cbox a b. (f x)\k \ (g x)\k \ i\k \ j\k" proof (rule ccontr) case goal1 then have *: "0 < (i\k - j\k) / 3" @@ -5010,7 +5130,7 @@ using rsum_component_le[OF p(1) goal1(3)] by (simp add: abs_real_def split: split_if_asm) qed - let ?P = "\a b. s = {a..b}" + let ?P = "\a b. s = cbox a b" { presume "\ ?P \ ?thesis" then show ?thesis @@ -5033,7 +5153,7 @@ from this[OF assms(2)] this[OF assms(3)] guess B1 B2 . note B=this[rule_format] have "bounded (ball 0 B1 \ ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+ - from bounded_subset_closed_interval[OF this] guess a b by (elim exE) + from bounded_subset_cbox[OF this] guess a b by (elim exE) note ab = conjunctD2[OF this[unfolded Un_subset_iff]] guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this] guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this] @@ -5052,7 +5172,7 @@ qed lemma integral_component_le: - fixes g f :: "'a::ordered_euclidean_space \ 'b::euclidean_space" + fixes g f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "f integrable_on s" "g integrable_on s" and "\x\s. (f x)\k \ (g x)\k" @@ -5063,7 +5183,7 @@ done lemma has_integral_component_nonneg: - fixes f :: "'a::ordered_euclidean_space \ 'b::euclidean_space" + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) s" and "\x\s. 0 \ (f x)\k" @@ -5073,7 +5193,7 @@ by auto lemma integral_component_nonneg: - fixes f :: "'a::ordered_euclidean_space \ 'b::euclidean_space" + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "f integrable_on s" "\x\s. 0 \ (f x)\k" shows "0 \ (integral s f)\k" @@ -5083,7 +5203,7 @@ done lemma has_integral_component_neg: - fixes f :: "'a::ordered_euclidean_space \ 'b::ordered_euclidean_space" + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) s" and "\x\s. (f x)\k \ 0" @@ -5092,29 +5212,29 @@ by auto lemma has_integral_component_lbound: - fixes f :: "'a::ordered_euclidean_space \ 'b::ordered_euclidean_space" - assumes "(f has_integral i) {a..b}" - and "\x\{a..b}. B \ f(x)\k" + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "(f has_integral i) (cbox a b)" + and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" - shows "B * content {a..b} \ i\k" + shows "B * content (cbox a b) \ i\k" using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) by (auto simp add: field_simps) lemma has_integral_component_ubound: - fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "(f has_integral i) {a..b}" - and "\x\{a..b}. f x\k \ B" + fixes f::"'a::euclidean_space => 'b::euclidean_space" + assumes "(f has_integral i) (cbox a b)" + and "\x\cbox a b. f x\k \ B" and "k \ Basis" - shows "i\k \ B * content {a..b}" + shows "i\k \ B * content (cbox a b)" using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) by (auto simp add: field_simps) lemma integral_component_lbound: - fixes f :: "'a::ordered_euclidean_space \ 'b::ordered_euclidean_space" - assumes "f integrable_on {a..b}" - and "\x\{a..b}. B \ f(x)\k" + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "f integrable_on cbox a b" + and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" - shows "B * content {a..b} \ (integral({a..b}) f)\k" + shows "B * content (cbox a b) \ (integral(cbox a b) f)\k" apply (rule has_integral_component_lbound) using assms unfolding has_integral_integral @@ -5122,11 +5242,11 @@ done lemma integral_component_ubound: - fixes f :: "'a::ordered_euclidean_space \ 'b::ordered_euclidean_space" - assumes "f integrable_on {a..b}" - and "\x\{a..b}. f x\k \ B" + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "f integrable_on cbox a b" + and "\x\cbox a b. f x\k \ B" and "k \ Basis" - shows "(integral {a..b} f)\k \ B * content {a..b}" + shows "(integral (cbox a b) f)\k \ B * content (cbox a b)" apply (rule has_integral_component_ubound) using assms unfolding has_integral_integral @@ -5137,12 +5257,12 @@ subsection {* Uniform limit of integrable functions is integrable. *} lemma integrable_uniform_limit: - fixes f :: "'a::ordered_euclidean_space \ 'b::banach" - assumes "\e>0. \g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" - shows "f integrable_on {a..b}" + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "\e>0. \g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" + shows "f integrable_on cbox a b" proof - { - presume *: "content {a..b} > 0 \ ?thesis" + presume *: "content (cbox a b) > 0 \ ?thesis" show ?thesis apply cases apply (rule *) @@ -5152,7 +5272,7 @@ apply auto done } - assume as: "content {a..b} > 0" + assume as: "content (cbox a b) > 0" have *: "\P. \e>(0::real). P e \ \n::nat. P (inverse (real n + 1))" by auto from choice[OF *[OF assms]] guess g .. note g=conjunctD2[OF this[rule_format],rule_format] @@ -5163,7 +5283,7 @@ proof (rule, rule) fix e :: real assume "e>0" - then have "e / 4 / content {a..b} > 0" + then have "e / 4 / content (cbox a b) > 0" using as by (auto simp add: field_simps) then guess M apply - @@ -5205,12 +5325,12 @@ apply (rule order_trans) apply (rule rsum_diff_bound[OF p(1), where e="2 / real M"]) proof - show "2 / real M * content {a..b} \ e / 2" + show "2 / real M * content (cbox a b) \ e / 2" unfolding divide_inverse using M as by (auto simp add: field_simps) fix x - assume x: "x \ {a..b}" + assume x: "x \ cbox a b" have "norm (f x - g n x) + norm (f x - g m x) \ inverse (real n + 1) + inverse (real m + 1)" using g(1)[OF x, of n] g(1)[OF x, of m] by auto also have "\ \ inverse (real M) + inverse (real M)" @@ -5237,7 +5357,7 @@ case goal1 then have *: "e/3 > 0" by auto from LIMSEQ_D [OF s this] guess N1 .. note N1=this - from goal1 as have "e / 3 / content {a..b} > 0" + from goal1 as have "e / 3 / content (cbox a b) > 0" by (auto simp add: field_simps) from real_arch_invD[OF this] guess N2 by (elim exE conjE) note N2=this from i[of "N1 + N2",unfolded has_integral,rule_format,OF *] guess g' .. note g'=conjunctD2[OF this,rule_format] @@ -5261,7 +5381,7 @@ apply (rule g') proof (rule, rule) fix p - assume p: "p tagged_division_of {a..b} \ g' fine p" + assume p: "p tagged_division_of (cbox a b) \ g' fine p" note * = g'(2)[OF this] show "norm ((\(x, k)\p. content k *\<^sub>R f x) - s) < e" apply - @@ -5272,15 +5392,15 @@ apply (rule g) apply assumption proof - - have "content {a..b} < e / 3 * (real N2)" + have "content (cbox a b) < e / 3 * (real N2)" using N2 unfolding inverse_eq_divide using as by (auto simp add: field_simps) - then have "content {a..b} < e / 3 * (real (N1 + N2) + 1)" + then have "content (cbox a b) < e / 3 * (real (N1 + N2) + 1)" apply - apply (rule less_le_trans,assumption) using `e>0` apply auto done - then show "inverse (real (N1 + N2) + 1) * content {a..b} \ e / 3" + then show "inverse (real (N1 + N2) + 1) * content (cbox a b) \ e / 3" unfolding inverse_eq_divide by (auto simp add: field_simps) show "norm (i (N1 + N2) - s) < e / 3" @@ -5293,8 +5413,8 @@ subsection {* Negligible sets. *} -definition "negligible (s:: 'a::ordered_euclidean_space set) \ - (\a b. ((indicator s :: 'a\real) has_integral 0) {a..b})" +definition "negligible (s:: 'a::euclidean_space set) \ + (\a b. ((indicator s :: 'a\real) has_integral 0) (cbox a b))" subsection {* Negligibility of hyperplane. *} @@ -5316,11 +5436,11 @@ done lemma interval_doublesplit: - fixes a :: "'a::ordered_euclidean_space" + fixes a :: "'a::euclidean_space" assumes "k \ Basis" - shows "{a..b} \ {x . abs(x\k - c) \ (e::real)} = - {(\i\Basis. (if i = k then max (a\k) (c - e) else a\i) *\<^sub>R i) .. - (\i\Basis. (if i = k then min (b\k) (c + e) else b\i) *\<^sub>R i)}" + shows "cbox a b \ {x . abs(x\k - c) \ (e::real)} = + cbox (\i\Basis. (if i = k then max (a\k) (c - e) else a\i) *\<^sub>R i) + (\i\Basis. (if i = k then min (b\k) (c + e) else b\i) *\<^sub>R i)" proof - have *: "\x c e::real. abs(x - c) \ e \ x \ c - e \ x \ c + e" by auto @@ -5331,10 +5451,10 @@ qed lemma division_doublesplit: - fixes a :: "'a::ordered_euclidean_space" - assumes "p division_of {a..b}" + fixes a :: "'a::euclidean_space" + assumes "p division_of (cbox a b)" and k: "k \ Basis" - shows "{l \ {x. abs(x\k - c) \ e} |l. l \ p \ l \ {x. abs(x\k - c) \ e} \ {}} division_of ({a..b} \ {x. abs(x\k - c) \ e})" + shows "{l \ {x. abs(x\k - c) \ e} |l. l \ p \ l \ {x. abs(x\k - c) \ e} \ {}} division_of (cbox a b \ {x. abs(x\k - c) \ e})" proof - have *: "\x c. abs (x - c) \ e \ x \ c - e \ x \ c + e" by auto @@ -5366,11 +5486,11 @@ qed lemma content_doublesplit: - fixes a :: "'a::ordered_euclidean_space" + fixes a :: "'a::euclidean_space" assumes "0 < e" and k: "k \ Basis" - obtains d where "0 < d" and "content ({a..b} \ {x. abs(x\k - c) \ d}) < e" -proof (cases "content {a..b} = 0") + obtains d where "0 < d" and "content (cbox a b \ {x. abs(x\k - c) \ d}) < e" +proof (cases "content (cbox a b) = 0") case True show ?thesis apply (rule that[of 1]) @@ -5402,9 +5522,9 @@ proof (rule that[of d]) have *: "Basis = insert k (Basis - {k})" using k by auto - have **: "{a..b} \ {x. \x \ k - c\ \ d} \ {} \ - (\i\Basis - {k}. Sup ({a..b} \ {x. \x \ k - c\ \ d}) \ i - - Inf ({a..b} \ {x. \x \ k - c\ \ d}) \ i) = + have **: "cbox a b \ {x. \x \ k - c\ \ d} \ {} \ + (\i\Basis - {k}. interval_upperbound (cbox a b \ {x. \x \ k - c\ \ d}) \ i - + interval_lowerbound (cbox a b \ {x. \x \ k - c\ \ d}) \ i) = (\i\Basis - {k}. b\i - a\i)" apply (rule setprod_cong) apply (rule refl) @@ -5412,10 +5532,10 @@ apply (subst interval_bounds) defer apply (subst interval_bounds) - unfolding interval_eq_empty not_ex not_less - apply auto - done - show "content ({a..b} \ {x. \x \ k - c\ \ d}) < e" + unfolding box_eq_empty not_ex not_less + apply auto + done + show "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" apply cases unfolding content_def apply (subst if_P) @@ -5425,7 +5545,7 @@ apply (subst *) apply (subst setprod_insert) unfolding ** - unfolding interval_doublesplit[OF k] interval_eq_empty not_ex not_less + unfolding interval_doublesplit[OF k] box_eq_empty not_ex not_less prefer 3 apply (subst interval_bounds) defer @@ -5445,7 +5565,7 @@ qed lemma negligible_standard_hyperplane[intro]: - fixes k :: "'a::ordered_euclidean_space" + fixes k :: "'a::euclidean_space" assumes k: "k \ Basis" shows "negligible {x. x\k = c}" unfolding negligible_def has_integral @@ -5461,7 +5581,7 @@ apply (rule d) proof (rule, rule) fix p - assume p: "p tagged_division_of {a..b} \ (\x. ball x d) fine p" + assume p: "p tagged_division_of (cbox a b) \ (\x. ball x d) fine p" have *: "(\(x, ka)\p. content ka *\<^sub>R ?i x) = (\(x, ka)\p. content (ka \ {x. abs(x\k - c) \ d}) *\<^sub>R ?i x)" apply (rule setsum_cong2) @@ -5526,7 +5646,7 @@ apply (subst setsum_over_tagged_division_lemma[OF p[THEN conjunct1]]) proof - case goal1 - have "content ({u..v} \ {x. \x \ k - c\ \ d}) \ content {u..v}" + have "content (cbox u v \ {x. \x \ k - c\ \ d}) \ content (cbox u v)" unfolding interval_doublesplit[OF k] apply (rule content_subset) unfolding interval_doublesplit[symmetric,OF k] @@ -5565,15 +5685,15 @@ proof (rule,rule,rule,rule,rule,rule,rule,erule conjE) fix m n u v assume as: - "{m..n} \ snd ` p" "{u..v} \ snd ` p" - "{m..n} \ {u..v}" - "{m..n} \ {x. \x \ k - c\ \ d} = {u..v} \ {x. \x \ k - c\ \ d}" - have "({m..n} \ {x. \x \ k - c\ \ d}) \ ({u..v} \ {x. \x \ k - c\ \ d}) \ {m..n} \ {u..v}" + "cbox m n \ snd ` p" "cbox u v \ snd ` p" + "cbox m n \ cbox u v" + "cbox m n \ {x. \x \ k - c\ \ d} = cbox u v \ {x. \x \ k - c\ \ d}" + have "(cbox m n \ {x. \x \ k - c\ \ d}) \ (cbox u v \ {x. \x \ k - c\ \ d}) \ cbox m n \ cbox u v" by blast - note interior_mono[OF this, unfolded division_ofD(5)[OF p'' as(1-3)] interior_inter[of "{m..n}"]] - then have "interior ({m..n} \ {x. \x \ k - c\ \ d}) = {}" + note interior_mono[OF this, unfolded division_ofD(5)[OF p'' as(1-3)] interior_inter[of "cbox m n"]] + then have "interior (cbox m n \ {x. \x \ k - c\ \ d}) = {}" unfolding as Int_absorb by auto - then show "content ({m..n} \ {x. \x \ k - c\ \ d}) = 0" + then show "content (cbox m n \ {x. \x \ k - c\ \ d}) = 0" unfolding interval_doublesplit[OF k] content_eq_0_interior[symmetric] . qed qed @@ -5586,18 +5706,18 @@ subsection {* A technical lemma about "refinement" of division. *} lemma tagged_division_finer: - fixes p :: "('a::ordered_euclidean_space \ ('a::ordered_euclidean_space set)) set" - assumes "p tagged_division_of {a..b}" + fixes p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" + assumes "p tagged_division_of (cbox a b)" and "gauge d" - obtains q where "q tagged_division_of {a..b}" + obtains q where "q tagged_division_of (cbox a b)" and "d fine q" and "\(x,k) \ p. k \ d(x) \ (x,k) \ q" proof - - let ?P = "\p. p tagged_partial_division_of {a..b} \ gauge d \ + let ?P = "\p. p tagged_partial_division_of (cbox a b) \ gauge d \ (\q. q tagged_division_of (\{k. \x. (x,k) \ p}) \ d fine q \ (\(x,k) \ p. k \ d(x) \ (x,k) \ q))" { - have *: "finite p" "p tagged_partial_division_of {a..b}" + have *: "finite p" "p tagged_partial_division_of (cbox a b)" using assms(1) unfolding tagged_division_of_def by auto @@ -5610,7 +5730,7 @@ apply auto done } - fix p :: "('a::ordered_euclidean_space \ ('a::ordered_euclidean_space set)) set" + fix p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" assume as: "finite p" show "?P p" apply rule @@ -5641,7 +5761,7 @@ using p apply auto done - then have int: "interior {u..v} \ interior (\{k. \x. (x, k) \ p}) = {}" + then have int: "interior (cbox u v) \ interior (\{k. \x. (x, k) \ p}) = {}" apply (rule inter_interior_unions_intervals) apply (rule open_interior) apply (rule_tac[!] ballI) @@ -5659,10 +5779,10 @@ apply auto done show ?case - proof (cases "{u..v} \ d x") + proof (cases "cbox u v \ d x") case True then show ?thesis - apply (rule_tac x="{(x,{u..v})} \ q1" in exI) + apply (rule_tac x="{(x,cbox u v)} \ q1" in exI) apply rule unfolding * uv apply (rule tagged_division_union) @@ -5767,13 +5887,13 @@ qed auto lemma has_integral_negligible: - fixes f :: "'b::ordered_euclidean_space \ 'a::real_normed_vector" + fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes "negligible s" and "\x\(t - s). f x = 0" shows "(f has_integral 0) t" proof - - presume P: "\f::'b::ordered_euclidean_space \ 'a. - \a b. \x. x \ s \ f x = 0 \ (f has_integral 0) {a..b}" + presume P: "\f::'b::euclidean_space \ 'a. + \a b. \x. x \ s \ f x = 0 \ (f has_integral 0) (cbox a b)" let ?f = "(\x. if x \ t then f x else 0)" show ?thesis apply (rule_tac f="?f" in has_integral_eq) @@ -5785,7 +5905,7 @@ apply (subst if_P, assumption) unfolding if_not_P proof - - assume "\a b. t = {a..b}" + assume "\a b. t = cbox a b" then guess a b apply - by (erule exE)+ note t = this show "(?f has_integral 0) t" unfolding t @@ -5795,8 +5915,8 @@ apply auto done next - show "\e>0. \B>0. \a b. ball 0 B \ {a..b} \ - (\z. ((\x. if x \ t then ?f x else 0) has_integral z) {a..b} \ norm (z - 0) < e)" + show "\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ t then ?f x else 0) has_integral z) (cbox a b) \ norm (z - 0) < e)" apply safe apply (rule_tac x=1 in exI) apply rule @@ -5813,7 +5933,7 @@ fix f :: "'b \ 'a" fix a b :: 'b assume assm: "\x. x \ s \ f x = 0" - show "(f has_integral 0) {a..b}" + show "(f has_integral 0) (cbox a b)" unfolding has_integral proof safe case goal1 @@ -5833,7 +5953,7 @@ show "gauge (\x. d (nat \norm (f x)\) x)" using d(1) unfolding gauge_def by auto fix p - assume as: "p tagged_division_of {a..b}" "(\x. d (nat \norm (f x)\) x) fine p" + assume as: "p tagged_division_of (cbox a b)" "(\x. d (nat \norm (f x)\) x) fine p" let ?goal = "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e" { presume "p \ {} \ ?goal" @@ -5850,7 +5970,7 @@ using as as' apply auto done - have "\i. \q. q tagged_division_of {a..b} \ (d i) fine q \ (\(x, k)\p. k \ (d i) x \ (x, k) \ q)" + have "\i. \q. q tagged_division_of (cbox a b) \ (d i) fine q \ (\(x, k)\p. k \ (d i) x \ (x, k) \ q)" apply rule apply (rule tagged_division_finer[OF as(1) d(1)]) apply auto @@ -5977,7 +6097,7 @@ qed lemma has_integral_spike: - fixes f :: "'b::ordered_euclidean_space \ 'a::real_normed_vector" + fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes "negligible s" and "(\x\(t - s). g x = f x)" and "(f has_integral y) t" @@ -5987,14 +6107,14 @@ fix a b :: 'b fix f g :: "'b \ 'a" fix y :: 'a - assume as: "\x \ {a..b} - s. g x = f x" "(f has_integral y) {a..b}" - have "((\x. f x + (g x - f x)) has_integral (y + 0)) {a..b}" + assume as: "\x \ cbox a b - s. g x = f x" "(f has_integral y) (cbox a b)" + have "((\x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" apply (rule has_integral_add[OF as(2)]) apply (rule has_integral_negligible[OF assms(1)]) using as apply auto done - then have "(g has_integral y) {a..b}" + then have "(g has_integral y) (cbox a b)" by auto } note * = this show ?thesis @@ -6105,7 +6225,7 @@ lemma negligible_union_eq[simp]: "negligible (s \ t) \ negligible s \ negligible t" using negligible_union by auto -lemma negligible_sing[intro]: "negligible {a::'a::ordered_euclidean_space}" +lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] by auto lemma negligible_insert[simp]: "negligible (insert a s) \ negligible s" @@ -6129,7 +6249,7 @@ using assms by induct auto lemma negligible: - "negligible s \ (\t::('a::ordered_euclidean_space) set. ((indicator s::'a\real) has_integral 0) t)" + "negligible s \ (\t::('a::euclidean_space) set. ((indicator s::'a\real) has_integral 0) t)" apply safe defer apply (subst negligible_def) @@ -6195,11 +6315,11 @@ subsection {* In particular, the boundary of an interval is negligible. *} -lemma negligible_frontier_interval: "negligible({a::'a::ordered_euclidean_space..b} - box a b)" +lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" proof - let ?A = "\((\k. {x. x\k = a\k} \ {x::'a. x\k = b\k}) ` Basis)" - have "{a..b} - box a b \ ?A" - apply rule unfolding Diff_iff mem_interval + have "cbox a b - box a b \ ?A" + apply rule unfolding Diff_iff mem_box apply simp apply(erule conjE bexE)+ apply(rule_tac x=i in bexI) @@ -6215,8 +6335,8 @@ lemma has_integral_spike_interior: assumes "\x\box a b. g x = f x" - and "(f has_integral y) ({a..b})" - shows "(g has_integral y) {a..b}" + and "(f has_integral y) (cbox a b)" + shows "(g has_integral y) (cbox a b)" apply (rule has_integral_spike[OF negligible_frontier_interval _ assms(2)]) using assms(1) apply auto @@ -6224,7 +6344,7 @@ lemma has_integral_spike_interior_eq: assumes "\x\box a b. g x = f x" - shows "(f has_integral y) {a..b} \ (g has_integral y) {a..b}" + shows "(f has_integral y) (cbox a b) \ (g has_integral y) (cbox a b)" apply rule apply (rule_tac[!] has_integral_spike_interior) using assms @@ -6233,8 +6353,8 @@ lemma integrable_spike_interior: assumes "\x\box a b. g x = f x" - and "f integrable_on {a..b}" - shows "g integrable_on {a..b}" + and "f integrable_on cbox a b" + shows "g integrable_on cbox a b" using assms unfolding integrable_on_def using has_integral_spike_interior[OF assms(1)] @@ -6260,21 +6380,21 @@ lemma operative_division_and: assumes "operative op \ P" - and "d division_of {a..b}" - shows "(\i\d. P i) \ P {a..b}" + and "d division_of (cbox a b)" + shows "(\i\d. P i) \ P (cbox a b)" using operative_division[OF monoidal_and assms] division_of_finite[OF assms(2)] by auto lemma operative_approximable: - fixes f::"'b::ordered_euclidean_space \ 'a::banach" + fixes f::"'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative op \ (\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i)" unfolding operative_def neutral_and proof safe fix a b :: 'b { - assume "content {a..b} = 0" - then show "\g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" + assume "content (cbox a b) = 0" + then show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" apply (rule_tac x=f in exI) using assms apply (auto intro!:integrable_on_null) @@ -6283,21 +6403,21 @@ { fix c g fix k :: 'b - assume as: "\x\{a..b}. norm (f x - g x) \ e" "g integrable_on {a..b}" + assume as: "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" assume k: "k \ Basis" - show "\g. (\x\{a..b} \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on {a..b} \ {x. x \ k \ c}" - "\g. (\x\{a..b} \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on {a..b} \ {x. c \ x \ k}" + show "\g. (\x\cbox a b \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. x \ k \ c}" + "\g. (\x\cbox a b \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. c \ x \ k}" apply (rule_tac[!] x=g in exI) using as(1) integrable_split[OF as(2) k] apply auto done } fix c k g1 g2 - assume as: "\x\{a..b} \ {x. x \ k \ c}. norm (f x - g1 x) \ e" "g1 integrable_on {a..b} \ {x. x \ k \ c}" - "\x\{a..b} \ {x. c \ x \ k}. norm (f x - g2 x) \ e" "g2 integrable_on {a..b} \ {x. c \ x \ k}" + assume as: "\x\cbox a b \ {x. x \ k \ c}. norm (f x - g1 x) \ e" "g1 integrable_on cbox a b \ {x. x \ k \ c}" + "\x\cbox a b \ {x. c \ x \ k}. norm (f x - g2 x) \ e" "g2 integrable_on cbox a b \ {x. c \ x \ k}" assume k: "k \ Basis" let ?g = "\x. if x\k = c then f x else if x\k \ c then g1 x else g2 x" - show "\g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" + show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" apply (rule_tac x="?g" in exI) proof safe case goal1 @@ -6310,13 +6430,13 @@ done next case goal2 - presume "?g integrable_on {a..b} \ {x. x \ k \ c}" - and "?g integrable_on {a..b} \ {x. x \ k \ c}" + presume "?g integrable_on cbox a b \ {x. x \ k \ c}" + and "?g integrable_on cbox a b \ {x. x \ k \ c}" then guess h1 h2 unfolding integrable_on_def by auto from has_integral_split[OF this k] show ?case unfolding integrable_on_def by auto next - show "?g integrable_on {a..b} \ {x. x \ k \ c}" "?g integrable_on {a..b} \ {x. x \ k \ c}" + show "?g integrable_on cbox a b \ {x. x \ k \ c}" "?g integrable_on cbox a b \ {x. x \ k \ c}" apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) using k as(2,4) apply auto @@ -6325,11 +6445,11 @@ qed lemma approximable_on_division: - fixes f :: "'b::ordered_euclidean_space \ 'a::banach" + fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" - and "d division_of {a..b}" + and "d division_of (cbox a b)" and "\i\d. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" - obtains g where "\x\{a..b}. norm (f x - g x) \ e" "g integrable_on {a..b}" + obtains g where "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" proof - note * = operative_division[OF monoidal_and operative_approximable[OF assms(1)] assms(2)] note this[unfolded iterate_and[OF division_of_finite[OF assms(2)]]] @@ -6342,13 +6462,13 @@ qed lemma integrable_continuous: - fixes f :: "'b::ordered_euclidean_space \ 'a::banach" - assumes "continuous_on {a..b} f" - shows "f integrable_on {a..b}" + fixes f :: "'b::euclidean_space \ 'a::banach" + assumes "continuous_on (cbox a b) f" + shows "f integrable_on cbox a b" proof (rule integrable_uniform_limit, safe) fix e :: real assume e: "e > 0" - from compact_uniformly_continuous[OF assms compact_interval,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d .. + from compact_uniformly_continuous[OF assms compact_cbox,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d .. note d=conjunctD2[OF this,rule_format] from fine_division_exists[OF gauge_ball[OF d(1)], of a b] guess p . note p=this note p' = tagged_division_ofD[OF p(1)] @@ -6376,10 +6496,16 @@ from e have "e \ 0" by auto from approximable_on_division[OF this division_of_tagged_division[OF p(1)] *] guess g . - then show "\g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" + then show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" by auto qed +lemma integrable_continuous_real: + fixes f :: "real \ 'a::banach" + assumes "continuous_on {a .. b} f" + shows "f integrable_on {a .. b}" + by (metis assms box_real(2) integrable_continuous) + subsection {* Specialization of additivity to one dimension. *} @@ -6388,26 +6514,36 @@ and real_inner_1_right: "inner x 1 = x" by simp_all +lemma content_real_eq_0: "content {a .. b::real} = 0 \ a \ b" + by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) + +lemma interval_real_split: + "{a .. b::real} \ {x. x \ c} = {a .. min b c}" + "{a .. b} \ {x. c \ x} = {max a c .. b}" + apply (metis Int_atLeastAtMostL1 atMost_def) + apply (metis Int_atLeastAtMostL2 atLeast_def) + done + lemma operative_1_lt: assumes "monoidal opp" - shows "operative opp f \ ((\a b. b \ a \ f {a..b::real} = neutral opp) \ - (\a b c. a < c \ c < b \ opp (f {a..c}) (f {c..b}) = f {a..b}))" - apply (simp add: operative_def content_eq_0) + shows "operative opp f \ ((\a b. b \ a \ f {a .. b::real} = neutral opp) \ + (\a b c. a < c \ c < b \ opp (f {a .. c}) (f {c .. b}) = f {a .. b}))" + apply (simp add: operative_def content_real_eq_0) proof safe fix a b c :: real assume as: - "\a b c. f {a..b} = opp (f ({a..b} \ {x. x \ c})) (f ({a..b} \ {x. c \ x}))" + "\a b c. f {a..b} = opp (f ({a..b} \ {x. x \ c})) (f ({a..b} \ Collect (op \ c)))" "a < c" "c < b" - from this(2-) have "{a..b} \ {x. x \ c} = {a..c}" "{a..b} \ {x. x \ c} = {c..b}" - by auto + from this(2-) have "cbox a b \ {x. x \ c} = cbox a c" "cbox a b \ {x. x \ c} = cbox c b" + by (auto simp: mem_box) then show "opp (f {a..c}) (f {c..b}) = f {a..b}" unfolding as(1)[rule_format,of a b "c"] by auto next fix a b c :: real assume as: "\a b. b \ a \ f {a..b} = neutral opp" "\a b c. a < c \ c < b \ opp (f {a..c}) (f {c..b}) = f {a..b}" - show "f {a..b} = opp (f ({a..b} \ {x. x \ c})) (f ({a..b} \ {x. c \ x}))" + show " f {a..b} = opp (f ({a..b} \ {x. x \ c})) (f ({a..b} \ Collect (op \ c)))" proof (cases "c \ {a..b}") case False then have "c < a \ c > b" by auto @@ -6424,7 +6560,7 @@ done next assume "b < c" - then have *: "{a..b} \ {x. x \ c} = {a..b}" "{a..b} \ {x. c \ x} = {1..0}" + then have *: "{a..b} \ {x. x \ c} = {a..b}" "{a..b} \ {x. c \ x} = {1 .. 0}" by auto show ?thesis unfolding * @@ -6442,7 +6578,7 @@ have ***: "\P Q. (\i\Basis. (if i = 1 then P i else Q i) *\<^sub>R i) = (P 1::real)" by simp show ?thesis - unfolding interval_split[OF **, unfolded real_inner_1_right] unfolding *** * + unfolding interval_real_split unfolding * proof (cases "c = a \ c = b") case False then show "f {a..b} = opp (f {a..c}) (f {c..b})" @@ -6456,7 +6592,7 @@ then show "f {a..b} = opp (f {a..c}) (f {c..b})" proof assume *: "c = a" - then have "f {a..c} = neutral opp" + then have "f {a .. c} = neutral opp" apply - apply (rule as(1)[rule_format]) apply auto @@ -6465,7 +6601,7 @@ using assms unfolding * by auto next assume *: "c = b" - then have "f {c..b} = neutral opp" + then have "f {c .. b} = neutral opp" apply - apply (rule as(1)[rule_format]) apply auto @@ -6479,8 +6615,8 @@ lemma operative_1_le: assumes "monoidal opp" - shows "operative opp f \ ((\a b. b \ a \ f {a..b::real} = neutral opp) \ - (\a b c. a \ c \ c \ b \ opp (f {a..c}) (f {c..b}) = f {a..b}))" + shows "operative opp f \ ((\a b. b \ a \ f {a .. b::real} = neutral opp) \ + (\a b c. a \ c \ c \ b \ opp (f {a .. c}) (f {c .. b}) = f {a .. b}))" unfolding operative_1_lt[OF assms] proof safe fix a b c :: real @@ -6495,7 +6631,7 @@ done next fix a b c :: real - assume "\a b. b \ a \ f {a..b} = neutral opp" + assume "\a b. b \ a \ f {a .. b} = neutral opp" and "\a b c. a < c \ c < b \ opp (f {a..c}) (f {c..b}) = f {a..b}" and "a \ c" and "c \ b" @@ -6514,7 +6650,7 @@ then show ?thesis proof assume *: "c = a" - then have "f {a..c} = neutral opp" + then have "f {a .. c} = neutral opp" apply - apply (rule as(1)[rule_format]) apply auto @@ -6523,7 +6659,7 @@ using assms unfolding * by auto next assume *: "c = b" - then have "f {c..b} = neutral opp" + then have "f {c .. b} = neutral opp" apply - apply (rule as(1)[rule_format]) apply auto @@ -6543,13 +6679,15 @@ and "p tagged_division_of {a..b}" shows "setsum (\(x,k). f(Sup k) - f(Inf k)) p = f b - f a" proof - - let ?f = "(\k::(real) set. if k = {} then 0 else f(Sup k) - f(Inf k))" + let ?f = "(\k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))" have ***: "\i\Basis. a \ i \ b \ i" using assms by auto have *: "operative op + ?f" - unfolding operative_1_lt[OF monoidal_monoid] interval_eq_empty by auto - have **: "{a..b} \ {}" - using assms(1) by auto note operative_tagged_division[OF monoidal_monoid * assms(2)] + unfolding operative_1_lt[OF monoidal_monoid] box_eq_empty + by auto + have **: "cbox a b \ {}" + using assms(1) by auto + note operative_tagged_division[OF monoidal_monoid * assms(2)[simplified box_real[symmetric]]] note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],symmetric] show ?thesis unfolding * @@ -6566,10 +6704,10 @@ subsection {* A useful lemma allowing us to factor out the content size. *} lemma has_integral_factor_content: - "(f has_integral i) {a..b} \ - (\e>0. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a..b}))" -proof (cases "content {a..b} = 0") + "(f has_integral i) (cbox a b) \ + (\e>0. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content (cbox a b)))" +proof (cases "content (cbox a b) = 0") case True show ?thesis unfolding has_integral_null_eq[OF True] @@ -6590,7 +6728,7 @@ case False note F = this[unfolded content_lt_nz[symmetric]] let ?P = "\e opp. \d. gauge d \ - (\p. p tagged_division_of {a..b} \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" + (\p. p tagged_division_of (cbox a b) \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" show ?thesis apply (subst has_integral) proof safe @@ -6598,8 +6736,8 @@ assume e: "e > 0" { assume "\e>0. ?P e op <" - then show "?P (e * content {a..b}) op \" - apply (erule_tac x="e * content {a..b}" in allE) + then show "?P (e * content (cbox a b)) op \" + apply (erule_tac x="e * content (cbox a b)" in allE) apply (erule impE) defer apply (erule exE,rule_tac x=d in exI) @@ -6608,9 +6746,9 @@ done } { - assume "\e>0. ?P (e * content {a..b}) op \" + assume "\e>0. ?P (e * content (cbox a b)) op \" then show "?P e op <" - apply (erule_tac x="e / 2 / content {a..b}" in allE) + apply (erule_tac x="e / 2 / content (cbox a b)" in allE) apply (erule impE) defer apply (erule exE,rule_tac x=d in exI) @@ -6621,6 +6759,13 @@ qed qed +lemma has_integral_factor_content_real: + "(f has_integral i) {a .. b::real} \ + (\e>0. \d. gauge d \ (\p. p tagged_division_of {a .. b} \ d fine p \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a .. b} ))" + unfolding box_real[symmetric] + by (rule has_integral_factor_content) + subsection {* Fundamental theorem of calculus. *} @@ -6629,28 +6774,25 @@ assumes "a \ b" shows "Sup {a..b} = b" and "Inf {a..b} = a" - apply (rule_tac[!] interval_bounds) - using assms - apply auto - done + using assms by auto lemma fundamental_theorem_of_calculus: fixes f :: "real \ 'a::banach" assumes "a \ b" - and "\x\{a..b}. (f has_vector_derivative f' x) (at x within {a..b})" - shows "(f' has_integral (f b - f a)) {a..b}" - unfolding has_integral_factor_content + and "\x\{a .. b}. (f has_vector_derivative f' x) (at x within {a .. b})" + shows "(f' has_integral (f b - f a)) {a .. b}" + unfolding has_integral_factor_content box_real[symmetric] proof safe fix e :: real assume e: "e > 0" note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt] - have *: "\P Q. \x\{a..b}. P x \ (\e>0. \d>0. Q x e d) \ \x. \(d::real)>0. x\{a..b} \ Q x e d" + have *: "\P Q. \x\{a .. b}. P x \ (\e>0. \d>0. Q x e d) \ \x. \(d::real)>0. x\{a .. b} \ Q x e d" using e by blast note this[OF assm,unfolded gauge_existence_lemma] from choice[OF this,unfolded Ball_def[symmetric]] guess d .. note d=conjunctD2[OF this[rule_format],rule_format] - show "\d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ - norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b})" + show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b))" apply (rule_tac x="\x. ball x (d x)" in exI) apply safe apply (rule gauge_ball_dependent) @@ -6658,10 +6800,10 @@ apply (rule d(1)) proof - fix p - assume as: "p tagged_division_of {a..b}" "(\x. ball x (d x)) fine p" - show "norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b}" - unfolding content_real[OF assms(1)] additive_tagged_division_1[OF assms(1) as(1),of f,symmetric] - unfolding additive_tagged_division_1[OF assms(1) as(1),of "\x. x",symmetric] + assume as: "p tagged_division_of cbox a b" "(\x. ball x (d x)) fine p" + show "norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b)" + unfolding content_real[OF assms(1), simplified box_real[symmetric]] additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of f,symmetric] + unfolding additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of "\x. x",symmetric] unfolding setsum_right_distrib defer unfolding setsum_subtractf[symmetric] @@ -6699,7 +6841,9 @@ by (auto simp add: dist_real_def field_simps) finally show "norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (Sup k - Inf k)" - unfolding k interval_bounds_real[OF *] content_real[OF *] . + unfolding box_real k interval_bounds_real[OF *] content_real[OF *] + interval_upperbound_real interval_lowerbound_real + . qed qed qed @@ -6723,19 +6867,19 @@ subsection {* Only need trivial subintervals if the interval itself is trivial. *} lemma division_of_nontrivial: - fixes s :: "'a::ordered_euclidean_space set set" - assumes "s division_of {a..b}" - and "content {a..b} \ 0" - shows "{k. k \ s \ content k \ 0} division_of {a..b}" + fixes s :: "'a::euclidean_space set set" + assumes "s division_of (cbox a b)" + and "content (cbox a b) \ 0" + shows "{k. k \ s \ content k \ 0} division_of (cbox a b)" using assms(1) apply - proof (induct "card s" arbitrary: s rule: nat_less_induct) fix s::"'a set set" - assume assm: "s division_of {a..b}" + assume assm: "s division_of (cbox a b)" "\mx. m = card x \ - x division_of {a..b} \ {k \ x. content k \ 0} division_of {a..b}" + x division_of (cbox a b) \ {k \ x. content k \ 0} division_of (cbox a b)" note s = division_ofD[OF assm(1)] - let ?thesis = "{k \ s. content k \ 0} division_of {a..b}" + let ?thesis = "{k \ s. content k \ 0} division_of (cbox a b)" { presume *: "{k \ s. content k \ 0} \ s \ ?thesis" show ?thesis @@ -6763,8 +6907,6 @@ defer apply rule apply (drule DiffD1,drule s(4)) - apply safe - apply (rule closed_interval) using assm(1) apply auto done @@ -6778,23 +6920,23 @@ assume as: "x \ k" "e > 0" from k(2)[unfolded k content_eq_0] guess i .. then have i:"c\i = d\i" "i\Basis" - using s(3)[OF k(1),unfolded k] unfolding interval_ne_empty by auto + using s(3)[OF k(1),unfolded k] unfolding box_ne_empty by auto then have xi: "x\i = d\i" - using as unfolding k mem_interval by (metis antisym) + using as unfolding k mem_box by (metis antisym) def y \ "\j\Basis. (if j = i then if c\i \ (a\i + b\i) / 2 then c\i + min e (b\i - c\i) / 2 else c\i - min e (c\i - a\i) / 2 else x\j) *\<^sub>R j" show "\x'\\(s - {k}). x' \ x \ dist x' x < e" apply (rule_tac x=y in bexI) proof - have "d \ {c..d}" + have "d \ cbox c d" using s(3)[OF k(1)] - unfolding k interval_eq_empty mem_interval + unfolding k box_eq_empty mem_box by (fastforce simp add: not_less) - then have "d \ {a..b}" + then have "d \ cbox a b" using s(2)[OF k(1)] unfolding k by auto - note di = this[unfolded mem_interval,THEN bspec[where x=i]] + note di = this[unfolded mem_box,THEN bspec[where x=i]] then have xyi: "y\i \ x\i" unfolding y_def i xi using as(2) assms(2)[unfolded content_eq_0] i(2) @@ -6818,7 +6960,7 @@ then show "dist y x < e" unfolding dist_norm by auto have "y \ k" - unfolding k mem_interval + unfolding k mem_box apply rule apply (erule_tac x=i in ballE) using xyi k i xi @@ -6827,15 +6969,15 @@ moreover have "y \ \s" using set_rev_mp[OF as(1) s(2)[OF k(1)]] as(2) di i - unfolding s mem_interval y_def + unfolding s mem_box y_def by (auto simp: field_simps elim!: ballE[of _ _ i]) ultimately show "y \ \(s - {k})" by auto qed qed - then have "\(s - {k}) = {a..b}" + then have "\(s - {k}) = cbox a b" unfolding s(6)[symmetric] by auto - then have "{ka \ s - {k}. content ka \ 0} division_of {a..b}" + then have "{ka \ s - {k}. content ka \ 0} division_of (cbox a b)" apply - apply (rule assm(2)[rule_format,OF card refl]) apply (rule division_ofI) @@ -6854,7 +6996,7 @@ subsection {* Integrability on subintervals. *} lemma operative_integrable: - fixes f :: "'b::ordered_euclidean_space \ 'a::banach" + fixes f :: "'b::euclidean_space \ 'a::banach" shows "operative op \ (\i. f integrable_on i)" unfolding operative_def neutral_and apply safe @@ -6866,17 +7008,24 @@ by (auto intro!: has_integral_split) lemma integrable_subinterval: - fixes f :: "'b::ordered_euclidean_space \ 'a::banach" - assumes "f integrable_on {a..b}" - and "{c..d} \ {a..b}" - shows "f integrable_on {c..d}" - apply (cases "{c..d} = {}") + fixes f :: "'b::euclidean_space \ 'a::banach" + assumes "f integrable_on cbox a b" + and "cbox c d \ cbox a b" + shows "f integrable_on cbox c d" + apply (cases "cbox c d = {}") defer apply (rule partial_division_extend_1[OF assms(2)],assumption) using operative_division_and[OF operative_integrable,symmetric,of _ _ _ f] assms(1) apply auto done +lemma integrable_subinterval_real: + fixes f :: "real \ 'a::banach" + assumes "f integrable_on {a .. b}" + and "{c .. d} \ {a .. b}" + shows "f integrable_on {c .. d}" + by (metis assms(1) assms(2) box_real(2) integrable_subinterval) + subsection {* Combining adjacent intervals in 1 dimension. *} @@ -6884,14 +7033,14 @@ fixes a b c :: real assumes "a \ c" and "c \ b" - and "(f has_integral i) {a..c}" - and "(f has_integral (j::'a::banach)) {c..b}" - shows "(f has_integral (i + j)) {a..b}" + and "(f has_integral i) {a .. c}" + and "(f has_integral (j::'a::banach)) {c .. b}" + shows "(f has_integral (i + j)) {a .. b}" proof - note operative_integral[of f, unfolded operative_1_le[OF monoidal_lifted[OF monoidal_monoid]]] note conjunctD2[OF this,rule_format] note * = this(2)[OF conjI[OF assms(1-2)],unfolded if_P[OF assms(3)]] - then have "f integrable_on {a..b}" + then have "f integrable_on cbox a b" apply - apply (rule ccontr) apply (subst(asm) if_P) @@ -6918,22 +7067,20 @@ fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" - and "f integrable_on {a..b}" - shows "integral {a..c} f + integral {c..b} f = integral {a..b} f" + and "f integrable_on {a .. b}" + shows "integral {a .. c} f + integral {c .. b} f = integral {a .. b} f" apply (rule integral_unique[symmetric]) apply (rule has_integral_combine[OF assms(1-2)]) - apply (rule_tac[!] integrable_integral integrable_subinterval[OF assms(3)])+ - using assms(1-2) - apply auto - done + apply (metis assms(2) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel2 monoid_add_class.add.left_neutral) + by (metis assms(1) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel1 monoid_add_class.add.right_neutral) lemma integrable_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" - and "f integrable_on {a..c}" - and "f integrable_on {c..b}" - shows "f integrable_on {a..b}" + and "f integrable_on {a .. c}" + and "f integrable_on {c .. b}" + shows "f integrable_on {a .. b}" using assms unfolding integrable_on_def by (fastforce intro!:has_integral_combine) @@ -6942,13 +7089,13 @@ subsection {* Reduce integrability to "local" integrability. *} lemma integrable_on_little_subintervals: - fixes f :: "'b::ordered_euclidean_space \ 'a::banach" - assumes "\x\{a..b}. \d>0. \u v. x \ {u..v} \ {u..v} \ ball x d \ {u..v} \ {a..b} \ - f integrable_on {u..v}" - shows "f integrable_on {a..b}" -proof - - have "\x. \d. x\{a..b} \ d>0 \ (\u v. x \ {u..v} \ {u..v} \ ball x d \ {u..v} \ {a..b} \ - f integrable_on {u..v})" + fixes f :: "'b::euclidean_space \ 'a::banach" + assumes "\x\cbox a b. \d>0. \u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ + f integrable_on cbox u v" + shows "f integrable_on cbox a b" +proof - + have "\x. \d. x\cbox a b \ d>0 \ (\u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ + f integrable_on cbox u v)" using assms by auto note this[unfolded gauge_existence_lemma] from choice[OF this] guess d .. note d=this[rule_format] @@ -6978,35 +7125,41 @@ subsection {* Second FCT or existence of antiderivative. *} -lemma integrable_const[intro]: "(\x. c) integrable_on {a..b}" +lemma integrable_const[intro]: "(\x. c) integrable_on cbox a b" unfolding integrable_on_def apply rule apply (rule has_integral_const) done +lemma integrable_const_ivl[intro]: + fixes a::"'a::ordered_euclidean_space" + shows "(\x. c) integrable_on {a .. b}" + unfolding cbox_interval[symmetric] + by (rule integrable_const) + lemma integral_has_vector_derivative: fixes f :: "real \ 'a::banach" - assumes "continuous_on {a..b} f" - and "x \ {a..b}" - shows "((\u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})" + assumes "continuous_on {a .. b} f" + and "x \ {a .. b}" + shows "((\u. integral {a .. u} f) has_vector_derivative f(x)) (at x within {a .. b})" unfolding has_vector_derivative_def has_derivative_within_alt apply safe apply (rule bounded_linear_scaleR_left) proof - fix e :: real assume e: "e > 0" - note compact_uniformly_continuous[OF assms(1) compact_interval,unfolded uniformly_continuous_on_def] + note compact_uniformly_continuous[OF assms(1) compact_Icc,unfolded uniformly_continuous_on_def] from this[rule_format,OF e] guess d by (elim conjE exE) note d=this[rule_format] - let ?I = "\a b. integral {a..b} f" - show "\d>0. \y\{a..b}. norm (y - x) < d \ + let ?I = "\a b. integral {a .. b} f" + show "\d>0. \y\{a .. b}. norm (y - x) < d \ norm (?I a y - ?I a x - (y - x) *\<^sub>R f x) \ e * norm (y - x)" proof (rule, rule, rule d, safe) case goal1 show ?case proof (cases "y < x") case False - have "f integrable_on {a..y}" - apply (rule integrable_subinterval,rule integrable_continuous) + have "f integrable_on {a .. y}" + apply (rule integrable_subinterval_real,rule integrable_continuous_real) apply (rule assms) unfolding not_less using assms(2) goal1 @@ -7021,34 +7174,28 @@ using assms(2) goal1 apply auto done - have **: "norm (y - x) = content {x..y}" - apply (subst content_real) - using False - unfolding not_less - apply auto - done + have **: "norm (y - x) = content {x .. y}" + using False by (auto simp: content_real) show ?thesis unfolding ** - apply (rule has_integral_bound[where f="(\u. f u - f x)"]) + apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) unfolding * - unfolding o_def defer apply (rule has_integral_sub) apply (rule integrable_integral) - apply (rule integrable_subinterval) - apply (rule integrable_continuous) + apply (rule integrable_subinterval_real) + apply (rule integrable_continuous_real) apply (rule assms)+ proof - - show "{x..y} \ {a..b}" + show "{x .. y} \ {a .. b}" using goal1 assms(2) by auto have *: "y - x = norm (y - x)" using False by auto - show "((\xa. f x) has_integral (y - x) *\<^sub>R f x) {x.. y}" + show "((\xa. f x) has_integral (y - x) *\<^sub>R f x) {x .. y}" apply (subst *) unfolding ** - apply auto - done - show "\xa\{x..y}. norm (f xa - f x) \ e" + by auto + show "\xa\{x .. y}. norm (f xa - f x) \ e" apply safe apply (rule less_imp_le) apply (rule d(2)[unfolded dist_norm]) @@ -7059,8 +7206,9 @@ qed (insert e, auto) next case True - have "f integrable_on {a..x}" + have "f integrable_on cbox a x" apply (rule integrable_subinterval,rule integrable_continuous) + unfolding box_real apply (rule assms)+ unfolding not_less using assms(2) goal1 @@ -7073,7 +7221,7 @@ using True using assms(2) goal1 apply auto done - have **: "norm (y - x) = content {y..x}" + have **: "norm (y - x) = content {y .. x}" apply (subst content_real) using True unfolding not_less @@ -7084,7 +7232,7 @@ show ?thesis apply (subst ***) unfolding norm_minus_cancel ** - apply (rule has_integral_bound[where f="(\u. f u - f x)"]) + apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) unfolding * unfolding o_def defer @@ -7092,19 +7240,19 @@ apply (subst minus_minus[symmetric]) unfolding minus_minus apply (rule integrable_integral) - apply (rule integrable_subinterval,rule integrable_continuous) + apply (rule integrable_subinterval_real,rule integrable_continuous_real) apply (rule assms)+ proof - - show "{y..x} \ {a..b}" + show "{y .. x} \ {a .. b}" using goal1 assms(2) by auto have *: "x - y = norm (y - x)" using True by auto - show "((\xa. f x) has_integral (x - y) *\<^sub>R f x) {y..x}" + show "((\xa. f x) has_integral (x - y) *\<^sub>R f x) {y .. x}" apply (subst *) unfolding ** apply auto done - show "\xa\{y..x}. norm (f xa - f x) \ e" + show "\xa\{y .. x}. norm (f xa - f x) \ e" apply safe apply (rule less_imp_le) apply (rule d(2)[unfolded dist_norm]) @@ -7119,8 +7267,8 @@ lemma antiderivative_continuous: fixes q b :: real - assumes "continuous_on {a..b} f" - obtains g where "\x\ {a..b}. (g has_vector_derivative (f x::_::banach)) (at x within {a..b})" + assumes "continuous_on {a .. b} f" + obtains g where "\x\{a .. b}. (g has_vector_derivative (f x::_::banach)) (at x within {a .. b})" apply (rule that) apply rule using integral_has_vector_derivative[OF assms] @@ -7132,15 +7280,15 @@ lemma antiderivative_integral_continuous: fixes f :: "real \ 'a::banach" - assumes "continuous_on {a..b} f" - obtains g where "\u\{a..b}. \v \ {a..b}. u \ v \ (f has_integral (g v - g u)) {u..v}" + assumes "continuous_on {a .. b} f" + obtains g where "\u\{a .. b}. \v \ {a .. b}. u \ v \ (f has_integral (g v - g u)) {u .. v}" proof - from antiderivative_continuous[OF assms] guess g . note g=this show ?thesis apply (rule that[of g]) proof safe case goal1 - have "\x\{u..v}. (g has_vector_derivative f x) (at x within {u..v})" + have "\x\cbox u v. (g has_vector_derivative f x) (at x within cbox u v)" apply rule apply (rule has_vector_derivative_within_subset) apply (rule g[rule_format]) @@ -7160,14 +7308,14 @@ and "\x. h(g x) = x" and "\x. g(h x) = x" and "\x. continuous (at x) g" - and "\u v. \w z. g ` {u..v} = {w..z}" - and "\u v. \w z. h ` {u..v} = {w..z}" - and "\u v. content(g ` {u..v}) = r * content {u..v}" - and "(f has_integral i) {a..b}" - shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` {a..b})" + and "\u v. \w z. g ` cbox u v = cbox w z" + and "\u v. \w z. h ` cbox u v = cbox w z" + and "\u v. content(g ` cbox u v) = r * content (cbox u v)" + and "(f has_integral i) (cbox a b)" + shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)" proof - { - presume *: "{a..b} \ {} \ ?thesis" + presume *: "cbox a b \ {} \ ?thesis" show ?thesis apply cases defer @@ -7178,7 +7326,7 @@ then show ?thesis unfolding goal1 assms(8)[unfolded goal1 has_integral_empty_eq] by auto qed } - assume "{a..b} \ {}" + assume "cbox a b \ {}" from assms(6)[rule_format,of a b] guess w z by (elim exE) note wz=this have inj: "inj g" "inj h" unfolding inj_on_def @@ -7210,7 +7358,7 @@ def d' \ "\x. {y. g y \ d (g x)}" have d': "\x. d' x = {y. g y \ (d (g x))}" unfolding d'_def .. - show "\d. gauge d \ (\p. p tagged_division_of h ` {a..b} \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" + show "\d. gauge d \ (\p. p tagged_division_of h ` cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" proof (rule_tac x=d' in exI, safe) show "gauge d'" using d(1) @@ -7218,9 +7366,9 @@ using continuous_open_preimage_univ[OF assms(4)] by auto fix p - assume as: "p tagged_division_of h ` {a..b}" "d' fine p" + assume as: "p tagged_division_of h ` cbox a b" "d' fine p" note p = tagged_division_ofD[OF as(1)] - have "(\(x, k). (g x, g ` k)) ` p tagged_division_of {a..b} \ d fine (\(x, k). (g x, g ` k)) ` p" + have "(\(x, k). (g x, g ` k)) ` p tagged_division_of (cbox a b) \ d fine (\(x, k). (g x, g ` k)) ` p" unfolding tagged_division_of proof safe show "finite ((\(x, k). (g x, g ` k)) ` p)" @@ -7231,12 +7379,12 @@ assume xk[intro]: "(x, k) \ p" show "g x \ g ` k" using p(2)[OF xk] by auto - show "\u v. g ` k = {u..v}" + show "\u v. g ` k = cbox u v" using p(4)[OF xk] using assms(5-6) by auto { fix y assume "y \ k" - then show "g y \ {a..b}" "g y \ {a..b}" + then show "g y \ cbox a b" "g y \ cbox a b" using p(3)[OF xk,unfolded subset_eq,rule_format,of "h (g y)"] using assms(2)[rule_format,of y] unfolding inj_image_mem_iff[OF inj(2)] @@ -7278,7 +7426,7 @@ } next fix x - assume "x \ {a..b}" + assume "x \ cbox a b" then have "h x \ \{k. \x. (x, k) \ p}" using p(6) by auto then guess X unfolding Union_iff .. note X=this @@ -7327,8 +7475,8 @@ subsection {* Special case of a basic affine transformation. *} lemma interval_image_affinity_interval: - "\u v. (\x. m *\<^sub>R (x::'a::ordered_euclidean_space) + c) ` {a..b} = {u..v}" - unfolding image_affinity_interval + "\u v. (\x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v" + unfolding image_affinity_cbox by auto lemma setprod_cong2: @@ -7339,12 +7487,12 @@ apply auto done -lemma content_image_affinity_interval: - "content((\x::'a::ordered_euclidean_space. m *\<^sub>R x + c) ` {a..b}) = - abs m ^ DIM('a) * content {a..b}" (is "?l = ?r") +lemma content_image_affinity_cbox: + "content((\x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) = + abs m ^ DIM('a) * content (cbox a b)" (is "?l = ?r") proof - { - presume *: "{a..b} \ {} \ ?thesis" + presume *: "cbox a b \ {} \ ?thesis" show ?thesis apply cases apply (rule *) @@ -7354,12 +7502,12 @@ apply auto done } - assume as: "{a..b} \ {}" + assume as: "cbox a b \ {}" show ?thesis proof (cases "m \ 0") case True - with as have "{m *\<^sub>R a + c..m *\<^sub>R b + c} \ {}" - unfolding interval_ne_empty + with as have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \ {}" + unfolding box_ne_empty apply (intro ballI) apply (erule_tac x=i in ballE) apply (auto simp: inner_simps intro!: mult_left_mono) @@ -7367,12 +7515,12 @@ moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b - a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis - by (simp add: image_affinity_interval True content_closed_interval' + by (simp add: image_affinity_cbox True content_cbox' setprod_timesf setprod_constant inner_diff_left) next case False - with as have "{m *\<^sub>R b + c..m *\<^sub>R a + c} \ {}" - unfolding interval_ne_empty + with as have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \ {}" + unfolding box_ne_empty apply (intro ballI) apply (erule_tac x=i in ballE) apply (auto simp: inner_simps intro!: mult_left_mono) @@ -7380,16 +7528,16 @@ moreover from False have *: "\i. (m *\<^sub>R a + c) \ i - (m *\<^sub>R b + c) \ i = (-m) *\<^sub>R (b - a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis using False - by (simp add: image_affinity_interval content_closed_interval' + by (simp add: image_affinity_cbox content_cbox' setprod_timesf[symmetric] setprod_constant[symmetric] inner_diff_left) qed qed lemma has_integral_affinity: - fixes a :: "'a::ordered_euclidean_space" - assumes "(f has_integral i) {a..b}" + fixes a :: "'a::euclidean_space" + assumes "(f has_integral i) (cbox a b)" and "m \ 0" - shows "((\x. f(m *\<^sub>R x + c)) has_integral ((1 / (abs(m) ^ DIM('a))) *\<^sub>R i)) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` {a..b})" + shows "((\x. f(m *\<^sub>R x + c)) has_integral ((1 / (abs(m) ^ DIM('a))) *\<^sub>R i)) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)" apply (rule has_integral_twiddle) apply safe apply (rule zero_less_power) @@ -7402,15 +7550,15 @@ apply (simp add: field_simps) apply (rule continuous_intros)+ apply (rule interval_image_affinity_interval)+ - apply (rule content_image_affinity_interval) + apply (rule content_image_affinity_cbox) using assms apply auto done lemma integrable_affinity: - assumes "f integrable_on {a..b}" + assumes "f integrable_on cbox a b" and "m \ 0" - shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` {a..b})" + shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)" using assms unfolding integrable_on_def apply safe @@ -7422,20 +7570,20 @@ subsection {* Special case of stretching coordinate axes separately. *} lemma image_stretch_interval: - "(\x. \k\Basis. (m k * (x\k)) *\<^sub>R k) ` {a..b::'a::ordered_euclidean_space} = - (if {a..b} = {} then {} else - {(\k\Basis. (min (m k * (a\k)) (m k * (b\k))) *\<^sub>R k)::'a .. - (\k\Basis. (max (m k * (a\k)) (m k * (b\k))) *\<^sub>R k)})" + "(\x. \k\Basis. (m k * (x\k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) = + (if (cbox a b) = {} then {} else + cbox (\k\Basis. (min (m k * (a\k)) (m k * (b\k))) *\<^sub>R k::'a) + (\k\Basis. (max (m k * (a\k)) (m k * (b\k))) *\<^sub>R k))" proof cases - assume *: "{a..b} \ {}" + assume *: "cbox a b \ {}" show ?thesis - unfolding interval_ne_empty if_not_P[OF *] - apply (simp add: interval image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) + unfolding box_ne_empty if_not_P[OF *] + apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) apply (subst choice_Basis_iff[symmetric]) proof (intro allI ball_cong refl) fix x i :: 'a assume "i \ Basis" with * have a_le_b: "a \ i \ b \ i" - unfolding interval_ne_empty by auto + unfolding box_ne_empty by auto show "(\xa. x \ i = m i * xa \ a \ i \ xa \ xa \ b \ i) \ min (m i * (a \ i)) (m i * (b \ i)) \ x \ i \ x \ i \ max (m i * (a \ i)) (m i * (b \ i))" proof (cases "m i = 0") @@ -7456,19 +7604,19 @@ qed simp lemma interval_image_stretch_interval: - "\u v. (\x. \k\Basis. (m k * (x\k))*\<^sub>R k) ` {a..b::'a::ordered_euclidean_space} = {u..v::'a}" + "\u v. (\x. \k\Basis. (m k * (x\k))*\<^sub>R k) ` cbox a (b::'a::euclidean_space) = cbox u (v::'a::euclidean_space)" unfolding image_stretch_interval by auto lemma content_image_stretch_interval: - "content ((\x::'a::ordered_euclidean_space. (\k\Basis. (m k * (x\k))*\<^sub>R k)::'a) ` {a..b}) = - abs (setprod m Basis) * content {a..b}" -proof (cases "{a..b} = {}") + "content ((\x::'a::euclidean_space. (\k\Basis. (m k * (x\k))*\<^sub>R k)::'a) ` cbox a b) = + abs (setprod m Basis) * content (cbox a b)" +proof (cases "cbox a b = {}") case True then show ?thesis unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto next case False - then have "(\x. (\k\Basis. (m k * (x\k))*\<^sub>R k)) ` {a..b} \ {}" + then have "(\x. (\k\Basis. (m k * (x\k))*\<^sub>R k)) ` cbox a b \ {}" by auto then show ?thesis using False @@ -7489,18 +7637,18 @@ apply - apply (erule disjE)+ unfolding min_def max_def - using False[unfolded interval_ne_empty,rule_format,of i] i + using False[unfolded box_ne_empty,rule_format,of i] i apply (auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos) done qed qed lemma has_integral_stretch: - fixes f :: "'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "(f has_integral i) {a..b}" + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes "(f has_integral i) (cbox a b)" and "\k\Basis. m k \ 0" shows "((\x. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) has_integral - ((1/(abs(setprod m Basis))) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` {a..b})" + ((1/(abs(setprod m Basis))) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` cbox a b)" apply (rule has_integral_twiddle[where f=f]) unfolding zero_less_abs_iff content_image_stretch_interval unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] @@ -7516,11 +7664,11 @@ qed auto lemma integrable_stretch: - fixes f :: "'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "f integrable_on {a..b}" + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes "f integrable_on cbox a b" and "\k\Basis. m k \ 0" shows "(\x::'a. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) integrable_on - ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` {a..b})" + ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` cbox a b)" using assms unfolding integrable_on_def apply - @@ -7534,35 +7682,50 @@ subsection {* even more special cases. *} lemma uminus_interval_vector[simp]: - fixes a b :: "'a::ordered_euclidean_space" - shows "uminus ` {a..b} = {-b..-a}" + fixes a b :: "'a::euclidean_space" + shows "uminus ` cbox a b = cbox (-b) (-a)" apply (rule set_eqI) apply rule defer unfolding image_iff apply (rule_tac x="-x" in bexI) - apply (auto simp add:minus_le_iff le_minus_iff eucl_le[where 'a='a]) + apply (auto simp add:minus_le_iff le_minus_iff mem_box) done lemma has_integral_reflect_lemma[intro]: - assumes "(f has_integral i) {a..b}" - shows "((\x. f(-x)) has_integral i) {-b..-a}" + assumes "(f has_integral i) (cbox a b)" + shows "((\x. f(-x)) has_integral i) (cbox (-b) (-a))" using has_integral_affinity[OF assms, of "-1" 0] by auto +lemma has_integral_reflect_lemma_real[intro]: + assumes "(f has_integral i) {a .. b::real}" + shows "((\x. f(-x)) has_integral i) {-b .. -a}" + using assms + unfolding box_real[symmetric] + by (rule has_integral_reflect_lemma) + lemma has_integral_reflect[simp]: - "((\x. f (-x)) has_integral i) {-b..-a} \ (f has_integral i) {a..b}" + "((\x. f (-x)) has_integral i) (cbox (-b) (-a)) \ (f has_integral i) (cbox a b)" apply rule apply (drule_tac[!] has_integral_reflect_lemma) apply auto done -lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on {-b..-a} \ f integrable_on {a..b}" +lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on cbox (-b) (-a) \ f integrable_on cbox a b" unfolding integrable_on_def by auto -lemma integral_reflect[simp]: "integral {-b..-a} (\x. f (-x)) = integral {a..b} f" +lemma integrable_reflect_real[simp]: "(\x. f(-x)) integrable_on {-b .. -a} \ f integrable_on {a .. b::real}" + unfolding box_real[symmetric] + by (rule integrable_reflect) + +lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\x. f (-x)) = integral (cbox a b) f" unfolding integral_def by auto +lemma integral_reflect_real[simp]: "integral {-b .. -a} (\x. f (-x)) = integral {a .. b::real} f" + unfolding box_real[symmetric] + by (rule integral_reflect) + subsection {* Stronger form of FCT; quite a tedious proof. *} @@ -7590,9 +7753,9 @@ lemma fundamental_theorem_of_calculus_interior: fixes f :: "real \ 'a::real_normed_vector" assumes "a \ b" - and "continuous_on {a..b} f" - and "\x\box a b. (f has_vector_derivative f'(x)) (at x)" - shows "(f' has_integral (f b - f a)) {a..b}" + and "continuous_on {a .. b} f" + and "\x\{a <..< b}. (f has_vector_derivative f'(x)) (at x)" + shows "(f' has_integral (f b - f a)) {a .. b}" proof - { presume *: "a < b \ ?thesis" @@ -7604,20 +7767,19 @@ case False then have "a = b" using assms(1) by auto - then have *: "{a .. b} = {b}" "f b - f a = 0" + then have *: "cbox a b = {b}" "f b - f a = 0" by (auto simp add: order_antisym) show ?thesis unfolding *(2) - apply (rule has_integral_null) unfolding content_eq_0 using * `a = b` by (auto simp: ex_in_conv) qed } assume ab: "a < b" - let ?P = "\e. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ - norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b})" - { presume "\e. e > 0 \ ?P e" then show ?thesis unfolding has_integral_factor_content by auto } + let ?P = "\e. \d. gauge d \ (\p. p tagged_division_of {a .. b} \ d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content {a .. b})" + { presume "\e. e > 0 \ ?P e" then show ?thesis unfolding has_integral_factor_content_real by auto } fix e :: real assume e: "e > 0" note assms(3)[unfolded has_vector_derivative_def has_derivative_at_alt ball_conj_distrib] @@ -7636,17 +7798,17 @@ from choice[OF this] guess d .. note conjunctD2[OF this[rule_format]] note d = this[rule_format] - have "bounded (f ` {a..b})" + have "bounded (f ` cbox a b)" apply (rule compact_imp_bounded compact_continuous_image)+ - using compact_interval assms + using compact_cbox assms apply auto done from this[unfolded bounded_pos] guess B .. note B = this[rule_format] - have "\da. 0 < da \ (\c. a \ c \ {a..c} \ {a..b} \ {a..c} \ ball a da \ - norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ (e * (b - a)) / 4)" + have "\da. 0 < da \ (\c. a \ c \ {a .. c} \ {a .. b} \ {a .. c} \ ball a da \ + norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \ (e * (b - a)) / 4)" proof - - have "a \ {a..b}" + have "a \ {a .. b}" using ab by auto note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] note * = this[unfolded continuous_within Lim_within,rule_format] @@ -7678,8 +7840,8 @@ apply (rule l k)+ proof - fix c - assume as: "a \ c" "{a..c} \ {a..b}" "{a..c} \ ball a (min k l)" - note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_interval] + assume as: "a \ c" "{a .. c} \ {a .. b}" "{a .. c} \ ball a (min k l)" + note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box] have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \ norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" by (rule norm_triangle_ineq4) also have "\ \ e * (b - a) / 8 + e * (b - a) / 8" @@ -7705,16 +7867,16 @@ apply (auto simp add: field_simps) done qed - finally show "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b - a) / 4" + finally show "norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \ e * (b - a) / 4" unfolding content_real[OF as(1)] by auto qed qed then guess da .. note da=conjunctD2[OF this,rule_format] - have "\db>0. \c\b. {c..b} \ {a..b} \ {c..b} \ ball b db \ - norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ (e * (b - a)) / 4" + have "\db>0. \c\b. {c .. b} \ {a .. b} \ {c .. b} \ ball b db \ + norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \ (e * (b - a)) / 4" proof - - have "b \ {a..b}" + have "b \ {a .. b}" using ab by auto note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" @@ -7746,7 +7908,7 @@ proof - fix c assume as: "c \ b" "{c..b} \ {a..b}" "{c..b} \ ball b (min k l)" - note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_interval] + note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box] have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \ norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" by (rule norm_triangle_ineq4) also have "\ \ e * (b - a) / 8 + e * (b - a) / 8" @@ -7773,7 +7935,7 @@ apply (auto simp add: field_simps) done qed - finally show "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b - a) / 4" + finally show "norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \ e * (b - a) / 4" unfolding content_real[OF as(1)] by auto qed qed @@ -7821,13 +7983,13 @@ "e * (Sup k - Inf k) / 2 < norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k)))" from p(4)[OF this(1)] guess u v by (elim exE) note k=this - then have "u \ v" and uv: "{u, v} \ {u..v}" + then have "u \ v" and uv: "{u, v} \ cbox u v" using p(2)[OF as(1)] by auto - note result = as(2)[unfolded k interval_bounds_real[OF this(1)] content_real[OF this(1)]] + note result = as(2)[unfolded k box_real interval_bounds_real[OF this(1)] content_real[OF this(1)]] assume as': "x \ a" "x \ b" then have "x \ box a b" - using p(2-3)[OF as(1)] by (auto simp: interval) + using p(2-3)[OF as(1)] by (auto simp: mem_box) note * = d(2)[OF this] have "norm ((v - u) *\<^sub>R f' (x) - (f (v) - f (u))) = norm ((f (u) - f (x) - (u - x) *\<^sub>R f' (x)) - (f (v) - f (x) - (v - x) *\<^sub>R f' (x)))" @@ -7871,6 +8033,7 @@ defer unfolding setsum_Un_disjoint[OF pA(2-),symmetric] pA(1)[symmetric] unfolding setsum_right_distrib[symmetric] + thm additive_tagged_division_1 apply (subst additive_tagged_division_1[OF _ as(1)]) apply (rule assms) proof - @@ -7878,7 +8041,7 @@ assume "(x, k) \ p \ {t. fst t \ {a, b}}" note xk=IntD1[OF this] from p(4)[OF this] guess u v by (elim exE) note uv=this - with p(2)[OF xk] have "{u..v} \ {}" + with p(2)[OF xk] have "cbox u v \ {}" by auto then show "0 \ e * ((Sup k) - (Inf k))" unfolding uv using e by (auto simp add: field_simps) @@ -7902,7 +8065,7 @@ using p(2)[OF xk(1)] by auto then have *: "u = v" using xk - unfolding uv content_eq_0 interval_eq_empty + unfolding uv content_eq_0 box_eq_empty by auto then show "content k *\<^sub>R (f' (x)) - (f ((Sup k)) - f ((Inf k))) = 0" using xk unfolding uv by auto @@ -7931,7 +8094,7 @@ apply (rule_tac[1-2] **) proof - let ?B = "\x. {t \ p. fst t = x \ content (snd t) \ 0}" - have pa: "\k. (a, k) \ p \ \v. k = {a .. v} \ a \ v" + have pa: "\k. (a, k) \ p \ \v. k = cbox a v \ a \ v" proof - case goal1 guess u v using p(4)[OF goal1] by (elim exE) note uv=this @@ -7939,7 +8102,7 @@ using p(2)[OF goal1] unfolding uv by auto have u: "u = a" proof (rule ccontr) - have "u \ {u..v}" + have "u \ cbox u v" using p(2-3)[OF goal1(1)] unfolding uv by auto have "u \ a" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto @@ -7955,7 +8118,7 @@ apply auto done qed - have pb: "\k. (b, k) \ p \ \v. k = {v .. b} \ b \ v" + have pb: "\k. (b, k) \ p \ \v. k = cbox v b \ b \ v" proof - case goal1 guess u v using p(4)[OF goal1] by (elim exE) note uv=this @@ -7963,7 +8126,7 @@ using p(2)[OF goal1] unfolding uv by auto have u: "v = b" proof (rule ccontr) - have "u \ {u..v}" + have "u \ cbox u v" using p(2-3)[OF goal1(1)] unfolding uv by auto have "v \ b" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto @@ -7989,14 +8152,14 @@ guess v using pa[OF k(1)] .. note v = conjunctD2[OF this] guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "min v v'" have "box a ?v \ k \ k'" - unfolding v v' by (auto simp add: interval) + unfolding v v' by (auto simp add: mem_box) note interior_mono[OF this,unfolded interior_inter] moreover have "(a + ?v)/2 \ box a ?v" using k(3-) unfolding v v' content_eq_0 not_le - by (auto simp add: interval) + by (auto simp add: mem_box) ultimately have "(a + ?v)/2 \ interior k \ interior k'" - unfolding interior_open[OF open_interval] by auto + unfolding interior_open[OF open_box] by auto then have *: "k = k'" apply - apply (rule ccontr) @@ -8020,12 +8183,12 @@ guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "max v v'" have "box ?v b \ k \ k'" - unfolding v v' by (auto simp: interval) + unfolding v v' by (auto simp: mem_box) note interior_mono[OF this,unfolded interior_inter] moreover have " ((b + ?v)/2) \ box ?v b" - using k(3-) unfolding v v' content_eq_0 not_le by (auto simp: interval) + using k(3-) unfolding v v' content_eq_0 not_le by (auto simp: mem_box) ultimately have " ((b + ?v)/2) \ interior k \ interior k'" - unfolding interior_open[OF open_interval] by auto + unfolding interior_open[OF open_box] by auto then have *: "k = k'" apply - apply (rule ccontr) @@ -8061,7 +8224,7 @@ apply (auto simp add:subset_eq dist_real_def v) done ultimately show ?case - unfolding v interval_bounds_real[OF v(2)] + unfolding v interval_bounds_real[OF v(2)] box_real apply - apply(rule da(2)[of "v"]) using goal1 fineD[OF as(2) goal1(1)] @@ -8093,7 +8256,7 @@ done ultimately show ?case unfolding v - unfolding interval_bounds_real[OF v(2)] + unfolding interval_bounds_real[OF v(2)] box_real apply - apply(rule db(2)[of "v"]) using goal1 fineD[OF as(2) goal1(1)] @@ -8115,9 +8278,9 @@ fixes f :: "real \ 'a::banach" assumes "finite s" and "a \ b" - and "continuous_on {a..b} f" - and "\x\box a b - s. (f has_vector_derivative f'(x)) (at x)" - shows "(f' has_integral (f b - f a)) {a..b}" + and "continuous_on {a .. b} f" + and "\x\{a <..< b} - s. (f has_vector_derivative f'(x)) (at x)" + shows "(f' has_integral (f b - f a)) {a .. b}" using assms proof (induct "card s" arbitrary: s a b) case 0 @@ -8154,7 +8317,7 @@ by auto case True then have "a \ c" "c \ b" - by (auto simp: interval) + by (auto simp: mem_box) then show ?thesis apply (subst *) apply (rule has_integral_combine) @@ -8163,18 +8326,18 @@ using Suc(3) unfolding cs proof - - show "continuous_on {a..c} f" "continuous_on {c..b} f" + show "continuous_on {a .. c} f" "continuous_on {c .. b} f" apply (rule_tac[!] continuous_on_subset[OF Suc(5)]) using True - apply (auto simp: interval) - done - let ?P = "\i j. \x\box i j - s'. (f has_vector_derivative f' x) (at x)" + apply (auto simp: mem_box) + done + let ?P = "\i j. \x\{i <..< j} - s'. (f has_vector_derivative f' x) (at x)" show "?P a c" "?P c b" apply safe apply (rule_tac[!] Suc(6)[rule_format]) using True unfolding cs - apply (auto simp: interval) + apply (auto simp: mem_box) done qed auto qed @@ -8184,22 +8347,22 @@ fixes f :: "real \ 'a::banach" assumes "finite s" and "a \ b" - and "continuous_on {a..b} f" - and "\x\{a..b} - s. (f has_vector_derivative f'(x)) (at x)" - shows "(f' has_integral (f(b) - f(a))) {a..b}" + and "continuous_on {a .. b} f" + and "\x\{a .. b} - s. (f has_vector_derivative f'(x)) (at x)" + shows "(f' has_integral (f b - f a)) {a .. b}" apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f']) using assms(4) - apply (auto simp: interval) + apply (auto simp: mem_box) done lemma indefinite_integral_continuous_left: fixes f:: "real \ 'a::banach" - assumes "f integrable_on {a..b}" + assumes "f integrable_on {a .. b}" and "a < c" and "c \ b" and "e > 0" obtains d where "d > 0" - and "\t. c - d < t \ t \ c \ norm (integral {a..c} f - integral {a..t} f) < e" + and "\t. c - d < t \ t \ c \ norm (integral {a .. c} f - integral {a .. t} f) < e" proof - have "\w>0. \t. c - w < t \ t < c \ norm (f c) * norm(c - t) < e / 3" proof (cases "f c = 0") @@ -8244,12 +8407,12 @@ have *: "e / 3 > 0" using assms by auto - have "f integrable_on {a..c}" - apply (rule integrable_subinterval[OF assms(1)]) + have "f integrable_on {a .. c}" + apply (rule integrable_subinterval_real[OF assms(1)]) using assms(2-3) apply auto done - from integrable_integral[OF this,unfolded has_integral,rule_format,OF *] guess d1 .. + from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d1 .. note d1 = conjunctD2[OF this,rule_format] def d \ "\x. ball x w \ d1 x" have "gauge d" @@ -8268,7 +8431,7 @@ using k(1) using assms(2) by auto fix t assume as: "c - ?d < t" "t \ c" - let ?thesis = "norm (integral {a..c} f - integral {a..t} f) < e" + let ?thesis = "norm (integral ({a .. c}) f - integral ({a .. t}) f) < e" { presume *: "t < c \ ?thesis" show ?thesis @@ -8282,17 +8445,17 @@ } assume "t < c" - have "f integrable_on {a..t}" - apply (rule integrable_subinterval[OF assms(1)]) + have "f integrable_on {a .. t}" + apply (rule integrable_subinterval_real[OF assms(1)]) using assms(2-3) as(2) apply auto done - from integrable_integral[OF this,unfolded has_integral,rule_format,OF *] guess d2 .. + from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d2 .. note d2 = conjunctD2[OF this,rule_format] def d3 \ "\x. if x \ t then d1 x \ d2 x else d1 x" have "gauge d3" using d2(1) d1(1) unfolding d3_def gauge_def by auto - from fine_division_exists[OF this, of a t] guess p . note p=this + from fine_division_exists_real[OF this, of a t] guess p . note p=this note p'=tagged_division_ofD[OF this(1)] have pt: "\(x,k)\p. x \ t" proof safe @@ -8308,14 +8471,14 @@ done note d2_fin = d2(2)[OF conjI[OF p(1) this]] - have *: "{a..c} \ {x. x \ 1 \ t} = {a..t}" "{a..c} \ {x. x \ 1 \ t} = {t..c}" + have *: "{a .. c} \ {x. x \ 1 \ t} = {a .. t}" "{a .. c} \ {x. x \ 1 \ t} = {t .. c}" using assms(2-3) as by (auto simp add: field_simps) - have "p \ {(c, {t..c})} tagged_division_of {a..c} \ d1 fine p \ {(c, {t..c})}" + have "p \ {(c, {t .. c})} tagged_division_of {a .. c} \ d1 fine p \ {(c, {t .. c})}" apply rule - apply (rule tagged_division_union_interval[of _ _ _ 1 "t"]) + apply (rule tagged_division_union_interval_real[of _ _ _ 1 "t"]) unfolding * apply (rule p) - apply (rule tagged_division_of_self) + apply (rule tagged_division_of_self_real) unfolding fine_def apply safe proof - @@ -8340,25 +8503,25 @@ by auto qed (insert as(2), auto) note d1_fin = d1(2)[OF this] - have *: "integral{a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + (\(x, k)\p. content k *\<^sub>R f x)) - - integral {a..c} f) + ((\(x, k)\p. content k *\<^sub>R f x) - integral {a..t} f) + (c - t) *\<^sub>R f c" + have *: "integral {a .. c} f - integral {a .. t} f = -(((c - t) *\<^sub>R f c + (\(x, k)\p. content k *\<^sub>R f x)) - + integral {a .. c} f) + ((\(x, k)\p. content k *\<^sub>R f x) - integral {a .. t} f) + (c - t) *\<^sub>R f c" "e = (e/3 + e/3) + e/3" by auto - have **: "(\(x, k)\p \ {(c, {t..c})}. content k *\<^sub>R f x) = + have **: "(\(x, k)\p \ {(c, {t .. c})}. content k *\<^sub>R f x) = (c - t) *\<^sub>R f c + (\(x, k)\p. content k *\<^sub>R f x)" proof - have **: "\x F. F \ {x} = insert x F" by auto - have "(c, {t..c}) \ p" + have "(c, cbox t c) \ p" proof safe case goal1 - from p'(2-3)[OF this] have "c \ {a..t}" + from p'(2-3)[OF this] have "c \ cbox a t" by auto then show False using `t < c` by auto qed then show ?thesis - unfolding ** + unfolding ** box_real apply - apply (subst setsum_insert) apply (rule p') @@ -8401,14 +8564,14 @@ lemma indefinite_integral_continuous_right: fixes f :: "real \ 'a::banach" - assumes "f integrable_on {a..b}" + assumes "f integrable_on {a .. b}" and "a \ c" and "c < b" and "e > 0" obtains d where "0 < d" - and "\t. c \ t \ t < c + d \ norm (integral {a..c} f - integral {a..t} f) < e" -proof - - have *: "(\x. f (- x)) integrable_on {- b..- a}" "- b < - c" "- c \ - a" + and "\t. c \ t \ t < c + d \ norm (integral {a .. c} f - integral {a .. t} f) < e" +proof - + have *: "(\x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \ - a" using assms by auto from indefinite_integral_continuous_left[OF * `e>0`] guess d . note d = this let ?d = "min d (b - c)" @@ -8420,8 +8583,8 @@ using d(1) assms(3) by auto fix t :: real assume as: "c \ t" "t < c + ?d" - have *: "integral {a..c} f = integral {a..b} f - integral {c..b} f" - "integral {a..t} f = integral {a..b} f - integral {t..b} f" + have *: "integral {a .. c} f = integral {a .. b} f - integral {c .. b} f" + "integral {a .. t} f = integral {a .. b} f - integral {t .. b} f" unfolding algebra_simps apply (rule_tac[!] integral_combine) using assms as @@ -8429,7 +8592,7 @@ done have "(- c) - d < (- t) \ - t \ - c" using as by auto note d(2)[rule_format,OF this] - then show "norm (integral {a..c} f - integral {a..t} f) < e" + then show "norm (integral {a .. c} f - integral {a .. t} f) < e" unfolding * unfolding integral_reflect apply (subst norm_minus_commute) @@ -8440,12 +8603,12 @@ lemma indefinite_integral_continuous: fixes f :: "real \ 'a::banach" - assumes "f integrable_on {a..b}" - shows "continuous_on {a..b} (\x. integral {a..x} f)" + assumes "f integrable_on {a .. b}" + shows "continuous_on {a .. b} (\x. integral {a .. x} f)" proof (unfold continuous_on_iff, safe) fix x e :: real - assume as: "x \ {a..b}" "e > 0" - let ?thesis = "\d>0. \x'\{a..b}. dist x' x < d \ dist (integral {a..x'} f) (integral {a..x} f) < e" + assume as: "x \ {a .. b}" "e > 0" + let ?thesis = "\d>0. \x'\{a .. b}. dist x' x < d \ dist (integral {a .. x'} f) (integral {a .. x} f) < e" { presume *: "a < b \ ?thesis" show ?thesis @@ -8454,12 +8617,11 @@ apply assumption proof - case goal1 - then have "{a..b} = {x}" + then have "cbox a b = {x}" using as(1) apply - apply (rule set_eqI) - unfolding atLeastAtMost_iff - apply (auto simp only: field_simps not_less) + apply auto done then show ?case using `e > 0` by auto qed @@ -8509,8 +8671,8 @@ show "0 < min d1 d2" using d1 d2 by auto fix y - assume "y \ {a..b}" and "dist y x < min d1 d2" - then show "dist (integral {a..y} f) (integral {a..x} f) < e" + assume "y \ {a .. b}" and "dist y x < min d1 d2" + then show "dist (integral {a .. y} f) (integral {a .. x} f) < e" apply (subst dist_commute) apply (cases "y < x") unfolding dist_norm @@ -8530,16 +8692,16 @@ lemma has_derivative_zero_unique_strong_interval: fixes f :: "real \ 'a::banach" assumes "finite k" - and "continuous_on {a..b} f" + and "continuous_on {a .. b} f" and "f a = y" - and "\x\({a..b} - k). (f has_derivative (\h. 0)) (at x within {a..b})" "x \ {a..b}" + and "\x\({a .. b} - k). (f has_derivative (\h. 0)) (at x within {a .. b})" "x \ {a .. b}" shows "f x = y" proof - have ab: "a \ b" using assms by auto have *: "a \ x" using assms(5) by auto - have "((\x. 0\'a) has_integral f x - f a) {a..x}" + have "((\x. 0\'a) has_integral f x - f a) {a .. x}" apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) *]) apply (rule continuous_on_subset[OF assms(2)]) defer @@ -8547,10 +8709,10 @@ unfolding has_vector_derivative_def apply (subst has_derivative_within_open[symmetric]) apply assumption - apply (rule open_interval) - apply (rule has_derivative_within_subset[where s="{a..b}"]) + apply (rule open_greaterThanLessThan) + apply (rule has_derivative_within_subset[where s="{a .. b}"]) using assms(4) assms(5) - apply (auto simp: interval) + apply (auto simp: mem_box) done note this[unfolded *] note has_integral_unique[OF has_integral_0 this] @@ -8562,7 +8724,7 @@ subsection {* Generalize a bit to any convex set. *} lemma has_derivative_zero_unique_strong_convex: - fixes f :: "'a::ordered_euclidean_space \ 'b::banach" + fixes f :: "'a::euclidean_space \ 'b::banach" assumes "convex s" and "finite k" and "continuous_on s f" @@ -8584,7 +8746,7 @@ } assume "x \ c" note conv = assms(1)[unfolded convex_alt,rule_format] - have as1: "continuous_on {0..1} (f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x))" + have as1: "continuous_on {0 ..1} (f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x))" apply (rule continuous_on_intros)+ apply (rule continuous_on_subset[OF assms(3)]) apply safe @@ -8623,14 +8785,14 @@ apply rule proof - fix t - assume as: "t \ {0..1} - {t. (1 - t) *\<^sub>R c + t *\<^sub>R x \ k}" + assume as: "t \ {0 .. 1} - {t. (1 - t) *\<^sub>R c + t *\<^sub>R x \ k}" have *: "c - t *\<^sub>R c + t *\<^sub>R x \ s - k" apply safe apply (rule conv[unfolded scaleR_simps]) using `x \ s` `c \ s` as by (auto simp add: algebra_simps) have "(f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x) has_derivative (\x. 0) \ (\z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) - (at t within {0..1})" + (at t within {0 .. 1})" apply (rule diff_chain_within) apply (rule has_derivative_add) unfolding scaleR_simps @@ -8643,7 +8805,7 @@ using `x \ s` `c \ s` apply auto done - then show "((\xa. f ((1 - xa) *\<^sub>R c + xa *\<^sub>R x)) has_derivative (\h. 0)) (at t within {0..1})" + then show "((\xa. f ((1 - xa) *\<^sub>R c + xa *\<^sub>R x)) has_derivative (\h. 0)) (at t within {0 .. 1})" unfolding o_def . qed auto then show ?thesis @@ -8655,7 +8817,7 @@ generalize to locally convex set with limpt-free set of exceptions. *} lemma has_derivative_zero_unique_strong_connected: - fixes f :: "'a::ordered_euclidean_space \ 'b::banach" + fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected s" and "open s" and "finite k" @@ -8711,14 +8873,14 @@ subsection {* Integrating characteristic function of an interval *} lemma has_integral_restrict_open_subinterval: - fixes f :: "'a::ordered_euclidean_space \ 'b::banach" - assumes "(f has_integral i) {c..d}" - and "{c..d} \ {a..b}" - shows "((\x. if x \ box c d then f x else 0) has_integral i) {a..b}" + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "(f has_integral i) (cbox c d)" + and "cbox c d \ cbox a b" + shows "((\x. if x \ box c d then f x else 0) has_integral i) (cbox a b)" proof - def g \ "\x. if x \box c d then f x else 0" { - presume *: "{c..d} \ {} \ ?thesis" + presume *: "cbox c d \ {} \ ?thesis" show ?thesis apply cases apply (rule *) @@ -8726,7 +8888,7 @@ proof - case goal1 then have *: "box c d = {}" - by (metis bot.extremum_uniqueI interval_open_subset_closed) + by (metis bot.extremum_uniqueI box_subset_cbox) show ?thesis using assms(1) unfolding * @@ -8734,14 +8896,14 @@ by auto qed } - assume "{c..d} \ {}" + assume "cbox c d \ {}" from partial_division_extend_1[OF assms(2) this] guess p . note p=this note mon = monoidal_lifted[OF monoidal_monoid] note operat = operative_division[OF this operative_integral p(1), symmetric] - let ?P = "(if g integrable_on {a..b} then Some (integral {a..b} g) else None) = Some i" + let ?P = "(if g integrable_on cbox a b then Some (integral (cbox a b) g) else None) = Some i" { presume "?P" - then have "g integrable_on {a..b} \ integral {a..b} g = i" + then have "g integrable_on cbox a b \ integral (cbox a b) g = i" apply - apply cases apply (subst(asm) if_P) @@ -8756,7 +8918,7 @@ note iterate_eq_neutral[OF mon,unfolded neutral_lifted[OF monoidal_monoid]] note * = this[unfolded neutral_add] - have iterate:"iterate (lifted op +) (p - {{c..d}}) + have iterate:"iterate (lifted op +) (p - {cbox c d}) (\i. if g integrable_on i then Some (integral i g) else None) = Some 0" proof (rule *, rule) case goal1 @@ -8764,22 +8926,22 @@ by auto note div = division_ofD(2-5)[OF p(1) this] from div(3) guess u v by (elim exE) note uv=this - have "interior x \ interior {c..d} = {}" + have "interior x \ interior (cbox c d) = {}" using div(4)[OF p(2)] goal1 by auto then have "(g has_integral 0) x" unfolding uv apply - apply (rule has_integral_spike_interior[where f="\x. 0"]) - unfolding g_def interior_closed_interval + unfolding g_def interior_cbox apply auto done then show ?case by auto qed - have *: "p = insert {c..d} (p - {{c..d}})" + have *: "p = insert (cbox c d) (p - {cbox c d})" using p by auto - have **: "g integrable_on {c..d}" + have **: "g integrable_on cbox c d" apply (rule integrable_spike_interior[where f=f]) unfolding g_def defer @@ -8788,7 +8950,7 @@ apply auto done moreover - have "integral {c..d} g = i" + have "integral (cbox c d) g = i" apply (rule has_integral_unique[OF _ assms(1)]) apply (rule has_integral_spike_interior[where f=g]) defer @@ -8811,28 +8973,28 @@ qed lemma has_integral_restrict_closed_subinterval: - fixes f :: "'a::ordered_euclidean_space \ 'b::banach" - assumes "(f has_integral i) {c..d}" - and "{c..d} \ {a..b}" - shows "((\x. if x \ {c..d} then f x else 0) has_integral i) {a..b}" + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "(f has_integral i) (cbox c d)" + and "cbox c d \ cbox a b" + shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)" proof - note has_integral_restrict_open_subinterval[OF assms] note * = has_integral_spike[OF negligible_frontier_interval _ this] show ?thesis apply (rule *[of c d]) - using interval_open_subset_closed[of c d] + using box_subset_cbox[of c d] apply auto done qed lemma has_integral_restrict_closed_subintervals_eq: - fixes f :: "'a::ordered_euclidean_space \ 'b::banach" - assumes "{c..d} \ {a..b}" - shows "((\x. if x \ {c..d} then f x else 0) has_integral i) {a..b} \ (f has_integral i) {c..d}" + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "cbox c d \ cbox a b" + shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b) \ (f has_integral i) (cbox c d)" (is "?l = ?r") -proof (cases "{c..d} = {}") +proof (cases "cbox c d = {}") case False - let ?g = "\x. if x \ {c..d} then f x else 0" + let ?g = "\x. if x \ cbox c d then f x else 0" show ?thesis apply rule defer @@ -8840,17 +9002,17 @@ apply assumption proof - assume ?l - then have "?g integrable_on {c..d}" + then have "?g integrable_on cbox c d" apply - apply (rule integrable_subinterval[OF _ assms]) apply auto done - then have *: "f integrable_on {c..d}" + then have *: "f integrable_on cbox c d" apply - apply (rule integrable_eq) apply auto done - then have "i = integral {c..d} f" + then have "i = integral (cbox c d) f" apply - apply (rule has_integral_unique) apply (rule `?l`) @@ -8866,14 +9028,14 @@ text {* Hence we can apply the limit process uniformly to all integrals. *} lemma has_integral': - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) s \ - (\e>0. \B>0. \a b. ball 0 B \ {a..b} \ - (\z. ((\x. if x \ s then f(x) else 0) has_integral z) {a..b} \ norm(z - i) < e))" + (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ s then f(x) else 0) has_integral z) (cbox a b) \ norm(z - i) < e))" (is "?l \ (\e>0. ?r e)") proof - { - presume *: "\a b. s = {a..b} \ ?thesis" + presume *: "\a b. s = cbox a b \ ?thesis" show ?thesis apply cases apply (rule *) @@ -8882,7 +9044,7 @@ apply auto done } - assume "\a b. s = {a..b}" + assume "\a b. s = cbox a b" then guess a b by (elim exE) note s=this from bounded_interval[of a b, THEN conjunct1, unfolded bounded_pos] guess B .. note B = conjunctD2[OF this,rule_format] show ?thesis @@ -8897,8 +9059,8 @@ apply (rule_tac x=i in exI) proof fix c d :: 'n - assume as: "ball 0 (B+1) \ {c..d}" - then show "((\x. if x \ s then f x else 0) has_integral i) {c..d}" + assume as: "ball 0 (B+1) \ cbox c d" + then show "((\x. if x \ s then f x else 0) has_integral i) (cbox c d)" unfolding s apply - apply (rule has_integral_restrict_closed_subinterval) @@ -8915,10 +9077,10 @@ from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format] def c \ "(\i\Basis. (- max B C) *\<^sub>R i)::'n" def d \ "(\i\Basis. max B C *\<^sub>R i)::'n" - have c_d: "{a..b} \ {c..d}" + have c_d: "cbox a b \ cbox c d" apply safe apply (drule B(2)) - unfolding mem_interval + unfolding mem_box proof case goal1 then show ?case @@ -8926,9 +9088,9 @@ unfolding c_def d_def by (auto simp add: field_simps setsum_negf) qed - have "ball 0 C \ {c..d}" + have "ball 0 C \ cbox c d" apply safe - unfolding mem_interval mem_ball dist_norm + unfolding mem_box mem_ball dist_norm proof case goal1 then show ?case @@ -8936,7 +9098,7 @@ unfolding c_def d_def by (auto simp: setsum_negf) qed - from C(2)[OF this] have "\y. (f has_integral y) {a..b}" + from C(2)[OF this] have "\y. (f has_integral y) (cbox a b)" unfolding has_integral_restrict_closed_subintervals_eq[OF c_d,symmetric] unfolding s by auto @@ -8950,10 +9112,10 @@ from as[rule_format,OF this] guess C .. note C=conjunctD2[OF this,rule_format] def c \ "(\i\Basis. (- max B C) *\<^sub>R i)::'n" def d \ "(\i\Basis. max B C *\<^sub>R i)::'n" - have c_d: "{a..b} \ {c..d}" + have c_d: "cbox a b \ cbox c d" apply safe apply (drule B(2)) - unfolding mem_interval + unfolding mem_box proof case goal1 then show ?case @@ -8961,9 +9123,9 @@ unfolding c_def d_def by (auto simp add: field_simps setsum_negf) qed - have "ball 0 C \ {c..d}" + have "ball 0 C \ cbox c d" apply safe - unfolding mem_interval mem_ball dist_norm + unfolding mem_box mem_ball dist_norm proof case goal1 then show ?case @@ -8990,7 +9152,7 @@ qed lemma has_integral_le: - fixes f :: "'n::ordered_euclidean_space \ real" + fixes f :: "'n::euclidean_space \ real" assumes "(f has_integral i) s" and "(g has_integral j) s" and "\x\s. f x \ g x" @@ -9000,7 +9162,7 @@ by auto lemma integral_le: - fixes f :: "'n::ordered_euclidean_space \ real" + fixes f :: "'n::euclidean_space \ real" assumes "f integrable_on s" and "g integrable_on s" and "\x\s. f x \ g x" @@ -9008,7 +9170,7 @@ by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)]) lemma has_integral_nonneg: - fixes f :: "'n::ordered_euclidean_space \ real" + fixes f :: "'n::euclidean_space \ real" assumes "(f has_integral i) s" and "\x\s. 0 \ f x" shows "0 \ i" @@ -9018,7 +9180,7 @@ by auto lemma integral_nonneg: - fixes f :: "'n::ordered_euclidean_space \ real" + fixes f :: "'n::euclidean_space \ real" assumes "f integrable_on s" and "\x\s. 0 \ f x" shows "0 \ integral s f" @@ -9042,12 +9204,12 @@ qed lemma has_integral_restrict_univ: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" shows "((\x. if x \ s then f x else 0) has_integral i) UNIV \ (f has_integral i) s" by auto lemma has_integral_on_superset: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\x. x \ s \ f x = 0" and "s \ t" and "(f has_integral i) s" @@ -9067,7 +9229,7 @@ qed lemma integrable_on_superset: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\x. x \ s \ f x = 0" and "s \ t" and "f integrable_on s" @@ -9077,7 +9239,7 @@ by (auto intro:has_integral_on_superset) lemma integral_restrict_univ[intro]: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" shows "f integrable_on s \ integral UNIV (\x. if x \ s then f x else 0) = integral s f" apply (rule integral_unique) unfolding has_integral_restrict_univ @@ -9085,12 +9247,12 @@ done lemma integrable_restrict_univ: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" shows "(\x. if x \ s then f x else 0) integrable_on UNIV \ f integrable_on s" unfolding integrable_on_def by auto -lemma negligible_on_intervals: "negligible s \ (\a b. negligible(s \ {a..b}))" (is "?l \ ?r") +lemma negligible_on_intervals: "negligible s \ (\a b. negligible(s \ cbox a b))" (is "?l \ ?r") proof assume ?r show ?l @@ -9106,7 +9268,7 @@ qed auto lemma has_integral_spike_set_eq: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((s - t) \ (t - s))" shows "(f has_integral y) s \ (f has_integral y) t" unfolding has_integral_restrict_univ[symmetric,of f] @@ -9114,7 +9276,7 @@ by (auto split: split_if_asm) lemma has_integral_spike_set[dest]: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((s - t) \ (t - s))" and "(f has_integral y) s" shows "(f has_integral y) t" @@ -9122,7 +9284,7 @@ by auto lemma integrable_spike_set[dest]: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((s - t) \ (t - s))" and "f integrable_on s" shows "f integrable_on t" @@ -9131,7 +9293,7 @@ unfolding has_integral_spike_set_eq[OF assms(1)] . lemma integrable_spike_set_eq: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((s - t) \ (t - s))" shows "f integrable_on s \ f integrable_on t" apply rule @@ -9176,7 +9338,7 @@ subsection {* More lemmas that are useful later *} lemma has_integral_subset_component_le: - fixes f :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes k: "k \ Basis" and as: "s \ t" "(f has_integral i) s" "(f has_integral j) t" "\x\t. 0 \ f(x)\k" shows "i\k \ j\k" @@ -9191,7 +9353,7 @@ qed lemma has_integral_subset_le: - fixes f :: "'n::ordered_euclidean_space \ real" + fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "(f has_integral i) s" and "(f has_integral j) t" @@ -9202,7 +9364,7 @@ by auto lemma integral_subset_component_le: - fixes f :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "k \ Basis" and "s \ t" and "f integrable_on s" @@ -9215,7 +9377,7 @@ done lemma integral_subset_le: - fixes f :: "'n::ordered_euclidean_space \ real" + fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "f integrable_on s" and "f integrable_on t" @@ -9227,10 +9389,10 @@ done lemma has_integral_alt': - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" - shows "(f has_integral i) s \ (\a b. (\x. if x \ s then f x else 0) integrable_on {a..b}) \ - (\e>0. \B>0. \a b. ball 0 B \ {a..b} \ - norm (integral {a..b} (\x. if x \ s then f x else 0) - i) < e)" + fixes f :: "'n::euclidean_space \ 'a::banach" + shows "(f has_integral i) s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ + (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e)" (is "?l = ?r") proof assume ?r @@ -9245,7 +9407,7 @@ apply rule apply (rule B) apply safe - apply (rule_tac x="integral {a..b} (\x. if x \ s then f x else 0)" in exI) + apply (rule_tac x="integral (cbox a b) (\x. if x \ s then f x else 0)" in exI) apply (drule B(2)[rule_format]) using integrable_integral[OF `?r`[THEN conjunct1,rule_format]] apply auto @@ -9260,22 +9422,22 @@ from as[OF zero_less_one] guess B .. note B=conjunctD2[OF this,rule_format] let ?a = "\i\Basis. min (a\i) (-B) *\<^sub>R i::'n" let ?b = "\i\Basis. max (b\i) B *\<^sub>R i::'n" - show "?f integrable_on {a..b}" + show "?f integrable_on cbox a b" proof (rule integrable_subinterval[of _ ?a ?b]) - have "ball 0 B \ {?a..?b}" + have "ball 0 B \ cbox ?a ?b" apply safe - unfolding mem_ball mem_interval dist_norm + unfolding mem_ball mem_box dist_norm proof case goal1 then show ?case using Basis_le_norm[of i x] by (auto simp add:field_simps) qed from B(2)[OF this] guess z .. note conjunct1[OF this] - then show "?f integrable_on {?a..?b}" + then show "?f integrable_on cbox ?a ?b" unfolding integrable_on_def by auto - show "{a..b} \ {?a..?b}" + show "cbox a b \ cbox ?a ?b" apply safe - unfolding mem_interval + unfolding mem_box apply rule apply (erule_tac x=i in ballE) apply auto @@ -9285,8 +9447,8 @@ fix e :: real assume "e > 0" from as[OF this] guess B .. note B=conjunctD2[OF this,rule_format] - show "\B>0. \a b. ball 0 B \ {a..b} \ - norm (integral {a..b} (\x. if x \ s then f x else 0) - i) < e" + show "\B>0. \a b. ball 0 B \ cbox a b \ + norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" apply rule apply rule apply (rule B) @@ -9304,12 +9466,12 @@ subsection {* Continuity of the integral (for a 1-dimensional interval). *} lemma integrable_alt: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" shows "f integrable_on s \ - (\a b. (\x. if x \ s then f x else 0) integrable_on {a..b}) \ - (\e>0. \B>0. \a b c d. ball 0 B \ {a..b} \ ball 0 B \ {c..d} \ - norm (integral {a..b} (\x. if x \ s then f x else 0) - - integral {c..d} (\x. if x \ s then f x else 0)) < e)" + (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ + (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ + norm (integral (cbox a b) (\x. if x \ s then f x else 0) - + integral (cbox c d) (\x. if x \ s then f x else 0)) < e)" (is "?l = ?r") proof assume ?l @@ -9340,7 +9502,7 @@ next assume ?r note as = conjunctD2[OF this,rule_format] - let ?cube = "\n. {(\i\Basis. - real n *\<^sub>R i)::'n .. (\i\Basis. real n *\<^sub>R i)} :: 'n set" + let ?cube = "\n. cbox (\i\Basis. - real n *\<^sub>R i::'n) (\i\Basis. real n *\<^sub>R i)" have "Cauchy (\n. integral (?cube n) (\x. if x \ s then f x else 0))" proof (unfold Cauchy_def, safe) case goal1 @@ -9351,7 +9513,7 @@ assume n: "n \ N" have "ball 0 B \ ?cube n" apply safe - unfolding mem_ball mem_interval dist_norm + unfolding mem_ball mem_box dist_norm proof case goal1 then show ?case @@ -9388,9 +9550,9 @@ show "0 < ?B" using B(1) by auto fix a b :: 'n - assume ab: "ball 0 ?B \ {a..b}" + assume ab: "ball 0 ?B \ cbox a b" from real_arch_simple[of ?B] guess n .. note n=this - show "norm (integral {a..b} (\x. if x \ s then f x else 0) - i) < e" + show "norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" apply (rule norm_triangle_half_l) apply (rule B(2)) defer @@ -9403,11 +9565,11 @@ assume x: "x \ ball 0 B" then have "x \ ball 0 ?B" by auto - then show "x \ {a..b}" + then show "x \ cbox a b" using ab by blast show "x \ ?cube n" using x - unfolding mem_interval mem_ball dist_norm + unfolding mem_box mem_ball dist_norm apply - proof case goal1 @@ -9422,18 +9584,18 @@ qed lemma integrable_altD: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" - shows "\a b. (\x. if x \ s then f x else 0) integrable_on {a..b}" - and "\e. e > 0 \ \B>0. \a b c d. ball 0 B \ {a..b} \ ball 0 B \ {c..d} \ - norm (integral {a..b} (\x. if x \ s then f x else 0) - integral {c..d} (\x. if x \ s then f x else 0)) < e" + shows "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" + and "\e. e > 0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ + norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" using assms[unfolded integrable_alt[of f]] by auto -lemma integrable_on_subinterval: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" +lemma integrable_on_subcbox: + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" - and "{a..b} \ s" - shows "f integrable_on {a..b}" + and "cbox a b \ s" + shows "f integrable_on cbox a b" apply (rule integrable_eq) defer apply (rule integrable_altD(1)[OF assms(1)]) @@ -9441,14 +9603,22 @@ apply auto done +lemma integrable_on_subinterval: + fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + assumes "f integrable_on s" + and "{a .. b} \ s" + shows "f integrable_on {a .. b}" + using integrable_on_subcbox[of f s a b] assms + by (simp add: cbox_interval) + subsection {* A straddling criterion for integrability *} lemma integrable_straddle_interval: - fixes f :: "'n::ordered_euclidean_space \ real" - assumes "\e>0. \g h i j. (g has_integral i) {a..b} \ (h has_integral j) {a..b} \ - norm (i - j) < e \ (\x\{a..b}. (g x) \ f x \ f x \ h x)" - shows "f integrable_on {a..b}" + fixes f :: "'n::euclidean_space \ real" + assumes "\e>0. \g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ + norm (i - j) < e \ (\x\cbox a b. (g x) \ f x \ f x \ h x)" + shows "f integrable_on cbox a b" proof (subst integrable_cauchy, safe) case goal1 then have e: "e/3 > 0" @@ -9522,12 +9692,12 @@ qed lemma integrable_straddle: - fixes f :: "'n::ordered_euclidean_space \ real" + fixes f :: "'n::euclidean_space \ real" assumes "\e>0. \g h i j. (g has_integral i) s \ (h has_integral j) s \ norm (i - j) < e \ (\x\s. g x \ f x \ f x \ h x)" shows "f integrable_on s" proof - - have "\a b. (\x. if x \ s then f x else 0) integrable_on {a..b}" + have "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" proof (rule integrable_straddle_interval, safe) case goal1 then have *: "e/4 > 0" @@ -9543,9 +9713,9 @@ from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format] def c \ "\i\Basis. min (a\i) (- (max B1 B2)) *\<^sub>R i::'n" def d \ "\i\Basis. max (b\i) (max B1 B2) *\<^sub>R i::'n" - have *: "ball 0 B1 \ {c..d}" "ball 0 B2 \ {c..d}" + have *: "ball 0 B1 \ cbox c d" "ball 0 B2 \ cbox c d" apply safe - unfolding mem_ball mem_interval dist_norm + unfolding mem_ball mem_box dist_norm apply (rule_tac[!] ballI) proof - case goal1 @@ -9564,8 +9734,8 @@ show ?case apply (rule_tac x="\x. if x \ s then g x else 0" in exI) apply (rule_tac x="\x. if x \ s then h x else 0" in exI) - apply (rule_tac x="integral {a..b} (\x. if x \ s then g x else 0)" in exI) - apply (rule_tac x="integral {a..b} (\x. if x \ s then h x else 0)" in exI) + apply (rule_tac x="integral (cbox a b) (\x. if x \ s then g x else 0)" in exI) + apply (rule_tac x="integral (cbox a b) (\x. if x \ s then h x else 0)" in exI) apply safe apply (rule_tac[1-2] integrable_integral,rule g) apply (rule h) @@ -9575,10 +9745,10 @@ (if x \ s then f x - g x else (0::real))" by auto note ** = abs_of_nonneg[OF integral_nonneg[OF integrable_sub, OF h g]] - show "norm (integral {a..b} (\x. if x \ s then h x else 0) - - integral {a..b} (\x. if x \ s then g x else 0)) \ - norm (integral {c..d} (\x. if x \ s then h x else 0) - - integral {c..d} (\x. if x \ s then g x else 0))" + show "norm (integral (cbox a b) (\x. if x \ s then h x else 0) - + integral (cbox a b) (\x. if x \ s then g x else 0)) \ + norm (integral (cbox c d) (\x. if x \ s then h x else 0) - + integral (cbox c d) (\x. if x \ s then g x else 0))" unfolding integral_sub[OF h g,symmetric] real_norm_def apply (subst **) defer @@ -9589,9 +9759,9 @@ apply (rule integrable_integral integrable_sub h g)+ proof safe fix x - assume "x \ {a..b}" - then show "x \ {c..d}" - unfolding mem_interval c_def d_def + assume "x \ cbox a b" + then show "x \ cbox c d" + unfolding mem_box c_def d_def apply - apply rule apply (erule_tac x=i in ballE) @@ -9626,7 +9796,7 @@ apply (rule B1) proof - fix a b c d :: 'n - assume as: "ball 0 (max B1 B2) \ {a..b}" "ball 0 (max B1 B2) \ {c..d}" + assume as: "ball 0 (max B1 B2) \ cbox a b" "ball 0 (max B1 B2) \ cbox c d" have **: "ball 0 B1 \ ball (0::'n) (max B1 B2)" "ball 0 B2 \ ball (0::'n) (max B1 B2)" by auto have *: "\ga gc ha hc fa fc::real. @@ -9634,7 +9804,7 @@ abs (hc - j) < e / 3 \ abs (i - j) < e / 3 \ ga \ fa \ fa \ ha \ gc \ fc \ fc \ hc \ abs (fa - fc) < e" by (simp add: abs_real_def split: split_if_asm) - show "norm (integral {a..b} (\x. if x \ s then f x else 0) - integral {c..d} + show "norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" unfolding real_norm_def apply (rule *) @@ -9669,7 +9839,7 @@ subsection {* Adding integrals over several sets *} lemma has_integral_union: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "(f has_integral i) s" and "(f has_integral j) t" and "negligible (s \ t)" @@ -9686,7 +9856,7 @@ qed lemma has_integral_unions: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "finite t" and "\s\t. (f has_integral (i s)) s" and "\s\t. \s'\t. s \ s' \ negligible (s \ s')" @@ -9733,7 +9903,7 @@ text {* In particular adding integrals over a division, maybe not of an interval. *} lemma has_integral_combine_division: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "d division_of s" and "\k\d. (f has_integral (i k)) k" shows "(f has_integral (setsum i d)) s" @@ -9750,9 +9920,9 @@ case goal1 from d(4)[OF this(1)] d(4)[OF this(2)] guess a c b d by (elim exE) note obt=this from d(5)[OF goal1] show ?case - unfolding obt interior_closed_interval + unfolding obt interior_cbox apply - - apply (rule negligible_subset[of "({a..b}-box a b) \ ({c..d}-box c d)"]) + apply (rule negligible_subset[of "(cbox a b-box a b) \ (cbox c d-box c d)"]) apply (rule negligible_union negligible_frontier_interval)+ apply auto done @@ -9760,7 +9930,7 @@ qed lemma integral_combine_division_bottomup: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "d division_of s" and "\k\d. f integrable_on k" shows "integral s f = setsum (\i. integral i f) d" @@ -9772,7 +9942,7 @@ done lemma has_integral_combine_division_topdown: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" and "d division_of k" and "k \ s" @@ -9785,7 +9955,7 @@ from division_ofD(2,4)[OF assms(2) this] show ?case apply safe - apply (rule integrable_on_subinterval) + apply (rule integrable_on_subcbox) apply (rule assms) using assms(3) apply auto @@ -9793,7 +9963,7 @@ qed lemma integral_combine_division_topdown: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" and "d division_of s" shows "integral s f = setsum (\i. integral i f) d" @@ -9804,7 +9974,7 @@ done lemma integrable_combine_division: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "d division_of s" and "\i\d. f integrable_on i" shows "f integrable_on s" @@ -9813,7 +9983,7 @@ by (metis has_integral_combine_division[OF assms(1)]) lemma integrable_on_subdivision: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "d division_of i" and "f integrable_on s" and "i \ s" @@ -9824,7 +9994,7 @@ note division_ofD(2,4)[OF assms(1) this] then show ?case apply safe - apply (rule integrable_on_subinterval[OF assms(2)]) + apply (rule integrable_on_subcbox[OF assms(2)]) using assms(3) apply auto done @@ -9834,7 +10004,7 @@ subsection {* Also tagged divisions *} lemma has_integral_combine_tagged_division: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "p tagged_division_of s" and "\(x,k) \ p. (f has_integral (i k)) k" shows "(f has_integral (setsum (\(x,k). i k) p)) s" @@ -9864,10 +10034,10 @@ qed lemma integral_combine_tagged_division_bottomup: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" - assumes "p tagged_division_of {a..b}" + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "p tagged_division_of (cbox a b)" and "\(x,k)\p. f integrable_on k" - shows "integral {a..b} f = setsum (\(x,k). integral k f) p" + shows "integral (cbox a b) f = setsum (\(x,k). integral k f) p" apply (rule integral_unique) apply (rule has_integral_combine_tagged_division[OF assms(1)]) using assms(2) @@ -9875,10 +10045,10 @@ done lemma has_integral_combine_tagged_division_topdown: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" - assumes "f integrable_on {a..b}" - and "p tagged_division_of {a..b}" - shows "(f has_integral (setsum (\(x,k). integral k f) p)) {a..b}" + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on cbox a b" + and "p tagged_division_of (cbox a b)" + shows "(f has_integral (setsum (\(x,k). integral k f) p)) (cbox a b)" apply (rule has_integral_combine_tagged_division[OF assms(2)]) proof safe case goal1 @@ -9888,10 +10058,10 @@ qed lemma integral_combine_tagged_division_topdown: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" - assumes "f integrable_on {a..b}" - and "p tagged_division_of {a..b}" - shows "integral {a..b} f = setsum (\(x,k). integral k f) p" + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on cbox a b" + and "p tagged_division_of (cbox a b)" + shows "integral (cbox a b) f = setsum (\(x,k). integral k f) p" apply (rule integral_unique) apply (rule has_integral_combine_tagged_division_topdown) using assms @@ -9902,13 +10072,13 @@ subsection {* Henstock's lemma *} lemma henstock_lemma_part1: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" - assumes "f integrable_on {a..b}" + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on cbox a b" and "e > 0" and "gauge d" - and "(\p. p tagged_division_of {a..b} \ d fine p \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p - integral({a..b}) f) < e)" - and p: "p tagged_partial_division_of {a..b}" "d fine p" + and "(\p. p tagged_division_of (cbox a b) \ d fine p \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p - integral(cbox a b) f) < e)" + and p: "p tagged_partial_division_of (cbox a b)" "d fine p" shows "norm (setsum (\(x,k). content k *\<^sub>R f x - integral k f) p) \ e" (is "?x \ e") proof - @@ -9916,7 +10086,7 @@ fix k :: real assume k: "k > 0" note p' = tagged_partial_division_ofD[OF p(1)] - have "\(snd ` p) \ {a..b}" + have "\(snd ` p) \ cbox a b" using p'(3) by fastforce note partial_division_of_tagged_division[OF p(1)] this from partial_division_extend_interval[OF this] guess q . note q=this and q' = division_ofD[OF this(2)] @@ -9938,7 +10108,7 @@ apply (rule k) apply auto done - have "f integrable_on {u..v}" + have "f integrable_on cbox u v" apply (rule integrable_subinterval[OF assms(1)]) using q'(2)[OF i] unfolding uv @@ -9958,7 +10128,7 @@ from bchoice[OF this] guess qq .. note qq=this[rule_format] let ?p = "p \ \(qq ` r)" - have "norm ((\(x, k)\?p. content k *\<^sub>R f x) - integral {a..b} f) < e" + have "norm ((\(x, k)\?p. content k *\<^sub>R f x) - integral (cbox a b) f) < e" apply (rule assms(4)[rule_format]) proof show "d fine ?p" @@ -10009,16 +10179,16 @@ apply auto done qed - moreover have "\(snd ` p) \ \r = {a..b}" and "{qq i |i. i \ r} = qq ` r" + moreover have "\(snd ` p) \ \r = cbox a b" and "{qq i |i. i \ r} = qq ` r" unfolding Union_Un_distrib[symmetric] r_def using q by auto - ultimately show "?p tagged_division_of {a..b}" + ultimately show "?p tagged_division_of (cbox a b)" by fastforce qed then have "norm ((\(x, k)\p. content k *\<^sub>R f x) + (\(x, k)\\(qq ` r). content k *\<^sub>R f x) - - integral {a..b} f) < e" + integral (cbox a b) f) < e" apply (subst setsum_Un_zero[symmetric]) apply (rule p') prefer 3 @@ -10045,7 +10215,7 @@ qed auto then have "norm ((\(x, k)\p. content k *\<^sub>R f x) + setsum (setsum (\(x, k). content k *\<^sub>R f x)) - (qq ` r) - integral {a..b} f) < e" + (qq ` r) - integral (cbox a b) f) < e" apply (subst (asm) setsum_UNION_zero) prefer 4 apply assumption @@ -10080,7 +10250,7 @@ qed (insert qq, auto) then have **: "norm ((\(x, k)\p. content k *\<^sub>R f x) + setsum (setsum (\(x, k). content k *\<^sub>R f x) \ qq) r - - integral {a..b} f) < e" + integral (cbox a b) f) < e" apply (subst (asm) setsum_reindex_nonzero) apply fact apply (rule setsum_0') @@ -10157,13 +10327,13 @@ qed lemma henstock_lemma_part2: - fixes f :: "'m::ordered_euclidean_space \ 'n::ordered_euclidean_space" - assumes "f integrable_on {a..b}" + fixes f :: "'m::euclidean_space \ 'n::euclidean_space" + assumes "f integrable_on cbox a b" and "e > 0" and "gauge d" - and "\p. p tagged_division_of {a..b} \ d fine p \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p - integral {a..b} f) < e" - and "p tagged_partial_division_of {a..b}" + and "\p. p tagged_division_of (cbox a b) \ d fine p \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p - integral (cbox a b) f) < e" + and "p tagged_partial_division_of (cbox a b)" and "d fine p" shows "setsum (\(x,k). norm (content k *\<^sub>R f x - integral k f)) p \ 2 * real (DIM('n)) * e" unfolding split_def @@ -10184,11 +10354,11 @@ done lemma henstock_lemma: - fixes f :: "'m::ordered_euclidean_space \ 'n::ordered_euclidean_space" - assumes "f integrable_on {a..b}" + fixes f :: "'m::euclidean_space \ 'n::euclidean_space" + assumes "f integrable_on cbox a b" and "e > 0" obtains d where "gauge d" - and "\p. p tagged_partial_division_of {a..b} \ d fine p \ + and "\p. p tagged_partial_division_of (cbox a b) \ d fine p \ setsum (\(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" proof - have *: "e / (2 * (real DIM('n) + 1)) > 0" @@ -10299,13 +10469,13 @@ subsection {* Monotone convergence (bounded interval first) *} lemma monotone_convergence_interval: - fixes f :: "nat \ 'n::ordered_euclidean_space \ real" - assumes "\k. (f k) integrable_on {a..b}" - and "\k. \x\{a..b}.(f k x) \ f (Suc k) x" - and "\x\{a..b}. ((\k. f k x) ---> g x) sequentially" - and "bounded {integral {a..b} (f k) | k . k \ UNIV}" - shows "g integrable_on {a..b} \ ((\k. integral ({a..b}) (f k)) ---> integral ({a..b}) g) sequentially" -proof (cases "content {a..b} = 0") + fixes f :: "nat \ 'n::euclidean_space \ real" + assumes "\k. (f k) integrable_on cbox a b" + and "\k. \x\cbox a b.(f k x) \ f (Suc k) x" + and "\x\cbox a b. ((\k. f k x) ---> g x) sequentially" + and "bounded {integral (cbox a b) (f k) | k . k \ UNIV}" + shows "g integrable_on cbox a b \ ((\k. integral (cbox a b) (f k)) ---> integral (cbox a b) g) sequentially" +proof (cases "content (cbox a b) = 0") case True show ?thesis using integrable_on_null[OF True] @@ -10314,7 +10484,7 @@ by auto next case False - have fg: "\x\{a..b}. \ k. (f k x) \ 1 \ (g x) \ 1" + have fg: "\x\cbox a b. \ k. (f k x) \ 1 \ (g x) \ 1" proof safe case goal1 note assms(3)[rule_format,OF this] @@ -10329,7 +10499,7 @@ apply auto done qed - have "\i. ((\k. integral ({a..b}) (f k)) ---> i) sequentially" + have "\i. ((\k. integral (cbox a b) (f k)) ---> i) sequentially" apply (rule bounded_increasing_convergent) defer apply rule @@ -10340,7 +10510,7 @@ apply auto done then guess i .. note i=this - have i': "\k. (integral({a..b}) (f k)) \ i\1" + have i': "\k. (integral(cbox a b) (f k)) \ i\1" apply (rule Lim_component_ge) apply (rule i) apply (rule trivial_limit_sequentially) @@ -10355,13 +10525,13 @@ apply auto done - have "(g has_integral i) {a..b}" + have "(g has_integral i) (cbox a b)" unfolding has_integral proof safe case goal1 note e=this - then have "\k. (\d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ - norm ((\(x, ka)\p. content ka *\<^sub>R f k x) - integral {a..b} (f k)) < e / 2 ^ (k + 2)))" + then have "\k. (\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ + norm ((\(x, ka)\p. content ka *\<^sub>R f k x) - integral (cbox a b) (f k)) < e / 2 ^ (k + 2)))" apply - apply rule apply (rule assms(1)[unfolded has_integral_integral has_integral,rule_format]) @@ -10370,7 +10540,7 @@ done from choice[OF this] guess c .. note c=conjunctD2[OF this[rule_format],rule_format] - have "\r. \k\r. 0 \ i\1 - (integral {a..b} (f k)) \ i\1 - (integral {a..b} (f k)) < e / 4" + have "\r. \k\r. 0 \ i\1 - (integral (cbox a b) (f k)) \ i\1 - (integral (cbox a b) (f k)) < e / 4" proof - case goal1 have "e/4 > 0" @@ -10388,11 +10558,11 @@ qed then guess r .. note r=conjunctD2[OF this[rule_format]] - have "\x\{a..b}. \n\r. \k\n. 0 \ (g x)\1 - (f k x)\1 \ - (g x)\1 - (f k x)\1 < e / (4 * content({a..b}))" + have "\x\cbox a b. \n\r. \k\n. 0 \ (g x)\1 - (f k x)\1 \ + (g x)\1 - (f k x)\1 < e / (4 * content(cbox a b))" proof case goal1 - have "e / (4 * content {a..b}) > 0" + have "e / (4 * content (cbox a b)) > 0" apply (rule divide_pos_pos) apply fact using False content_pos_le[of a b] @@ -10419,7 +10589,7 @@ using c(1) unfolding gauge_def d_def by auto next fix p - assume p: "p tagged_division_of {a..b}" "d fine p" + assume p: "p tagged_division_of (cbox a b)" "d fine p" note p'=tagged_division_ofD[OF p(1)] have "\a. \x\p. m (fst x) \ a" by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI) @@ -10440,7 +10610,7 @@ proof safe case goal1 show ?case - apply (rule order_trans[of _ "\(x, k)\p. content k * (e / (4 * content {a..b}))"]) + apply (rule order_trans[of _ "\(x, k)\p. content k * (e / (4 * content (cbox a b)))"]) unfolding setsum_subtractf[symmetric] apply (rule order_trans) apply (rule norm_setsum) @@ -10451,10 +10621,10 @@ proof - fix x k assume xk: "(x, k) \ p" - then have x: "x \ {a..b}" + then have x: "x \ cbox a b" using p'(2-3)[OF xk] by auto from p'(4)[OF xk] guess u v by (elim exE) note uv=this - show "norm (content k *\<^sub>R (g x - f (m x) x)) \ content k * (e / (4 * content {a..b}))" + show "norm (content k *\<^sub>R (g x - f (m x) x)) \ content k * (e / (4 * content (cbox a b)))" unfolding norm_scaleR uv unfolding abs_of_nonneg[OF content_pos_le] apply (rule mult_left_mono) @@ -10536,9 +10706,9 @@ unfolding split_paired_all split_conv apply (rule_tac[1-2] integral_le[OF ]) proof safe - show "0 \ i\1 - (integral {a..b} (f r))\1" + show "0 \ i\1 - (integral (cbox a b) (f r))\1" using r(1) by auto - show "i\1 - (integral {a..b} (f r))\1 < e / 4" + show "i\1 - (integral (cbox a b) (f r))\1 < e / 4" using r(2) by auto fix x k assume xk: "(x, k) \ p" @@ -10548,14 +10718,14 @@ and "f (m x) integrable_on k" and "f (m x) integrable_on k" unfolding uv - apply (rule_tac[!] integrable_on_subinterval[OF assms(1)[rule_format]]) + apply (rule_tac[!] integrable_on_subcbox[OF assms(1)[rule_format]]) using p'(3)[OF xk] unfolding uv apply auto done fix y assume "y \ k" - then have "y \ {a..b}" + then have "y \ cbox a b" using p'(3)[OF xk] by auto then have *: "\m. \n\m. f m y \ f n y" apply - @@ -10573,21 +10743,21 @@ qed qed note * = this - have "integral {a..b} g = i" + have "integral (cbox a b) g = i" by (rule integral_unique) (rule *) then show ?thesis using i * by auto qed lemma monotone_convergence_increasing: - fixes f :: "nat \ 'n::ordered_euclidean_space \ real" + fixes f :: "nat \ 'n::euclidean_space \ real" assumes "\k. (f k) integrable_on s" and "\k. \x\s. (f k x) \ (f (Suc k) x)" and "\x\s. ((\k. f k x) ---> g x) sequentially" and "bounded {integral s (f k)| k. True}" shows "g integrable_on s \ ((\k. integral s (f k)) ---> integral s g) sequentially" proof - - have lem: "\f::nat \ 'n::ordered_euclidean_space \ real. + have lem: "\f::nat \ 'n::euclidean_space \ real. \g s. \k.\x\s. 0 \ f k x \ \k. (f k) integrable_on s \ \k. \x\s. f k x \ f (Suc k) x \ \x\s. ((\k. f k x) ---> g x) sequentially \ bounded {integral s (f k)| k. True} \ @@ -10644,15 +10814,15 @@ have ifif: "\k t. (\x. if x \ t then if x \ s then f k x else 0 else 0) = (\x. if x \ t \ s then f k x else 0)" by (rule ext) auto - have int': "\k a b. f k integrable_on {a..b} \ s" + have int': "\k a b. f k integrable_on cbox a b \ s" apply (subst integrable_restrict_univ[symmetric]) apply (subst ifif[symmetric]) apply (subst integrable_restrict_univ) apply (rule int) done - have "\a b. (\x. if x \ s then g x else 0) integrable_on {a..b} \ - ((\k. integral {a..b} (\x. if x \ s then f k x else 0)) ---> - integral {a..b} (\x. if x \ s then g x else 0)) sequentially" + have "\a b. (\x. if x \ s then g x else 0) integrable_on cbox a b \ + ((\k. integral (cbox a b) (\x. if x \ s then f k x else 0)) ---> + integral (cbox a b) (\x. if x \ s then g x else 0)) sequentially" proof (rule monotone_convergence_interval, safe) case goal1 show ?case by (rule int) @@ -10673,7 +10843,7 @@ next case goal4 note * = integral_nonneg - have "\k. norm (integral {a..b} (\x. if x \ s then f k x else 0)) \ norm (integral s (f k))" + have "\k. norm (integral (cbox a b) (\x. if x \ s then f k x else 0)) \ norm (integral s (f k))" unfolding real_norm_def apply (subst abs_of_nonneg) apply (rule *[OF int]) @@ -10722,11 +10892,11 @@ apply safe proof - fix a b :: 'n - assume ab: "ball 0 B \ {a..b}" + assume ab: "ball 0 B \ cbox a b" from `e > 0` have "e/2 > 0" by auto from LIMSEQ_D [OF g(2)[of a b] this] guess M .. note M=this - have **: "norm (integral {a..b} (\x. if x \ s then f N x else 0) - i) < e/2" + have **: "norm (integral (cbox a b) (\x. if x \ s then f N x else 0) - i) < e/2" apply (rule norm_triangle_half_l) using B(2)[rule_format,OF ab] N[rule_format,of N] apply - @@ -10737,7 +10907,7 @@ have *: "\f1 f2 g. abs (f1 - i) < e / 2 \ abs (f2 - g) < e / 2 \ f1 \ f2 \ f2 \ i \ abs (g - i) < e" unfolding real_inner_1_right by arith - show "norm (integral {a..b} (\x. if x \ s then g x else 0) - i) < e" + show "norm (integral (cbox a b) (\x. if x \ s then g x else 0) - i) < e" unfolding real_norm_def apply (rule *[rule_format]) apply (rule **[unfolded real_norm_def]) @@ -10845,7 +11015,7 @@ qed lemma monotone_convergence_decreasing: - fixes f :: "nat \ 'n::ordered_euclidean_space \ real" + fixes f :: "nat \ 'n::euclidean_space \ real" assumes "\k. (f k) integrable_on s" and "\k. \x\s. f (Suc k) x \ f k x" and "\x\s. ((\k. f k x) ---> g x) sequentially" @@ -10910,12 +11080,12 @@ by auto (*lemma absolutely_integrable_on_trans[simp]: - fixes f::"'n::ordered_euclidean_space \ real" + fixes f::"'n::euclidean_space \ real" shows "(vec1 o f) absolutely_integrable_on s \ f absolutely_integrable_on s" unfolding absolutely_integrable_on_def o_def by auto*) lemma integral_norm_bound_integral: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" and "g integrable_on s" and "\x\s. norm (f x) \ g x" @@ -10945,8 +11115,8 @@ done qed note norm=this[rule_format] - have lem: "\f::'n \ 'a. \g a b. f integrable_on {a..b} \ g integrable_on {a..b} \ - \x\{a..b}. norm (f x) \ g x \ norm (integral({a..b}) f) \ integral {a..b} g" + have lem: "\f::'n \ 'a. \g a b. f integrable_on cbox a b \ g integrable_on cbox a b \ + \x\cbox a b. norm (f x) \ g x \ norm (integral(cbox a b) f) \ integral (cbox a b) g" proof (rule *[rule_format]) case goal1 then have *: "e/2 > 0" @@ -10992,13 +11162,13 @@ guess B1 .. note B1=conjunctD2[OF this[rule_format],rule_format] from integrable_integral[OF assms(2),unfolded has_integral'[of g],rule_format,OF e] guess B2 .. note B2=conjunctD2[OF this[rule_format],rule_format] - from bounded_subset_closed_interval[OF bounded_ball, of "0::'n" "max B1 B2"] + from bounded_subset_cbox[OF bounded_ball, of "0::'n" "max B1 B2"] guess a b by (elim exE) note ab=this[unfolded ball_max_Un] - have "ball 0 B1 \ {a..b}" + have "ball 0 B1 \ cbox a b" using ab by auto from B1(2)[OF this] guess z .. note z=conjunctD2[OF this] - have "ball 0 B2 \ {a..b}" + have "ball 0 B2 \ cbox a b" using ab by auto from B2(2)[OF this] guess w .. note w=conjunctD2[OF this] @@ -11018,8 +11188,8 @@ qed lemma integral_norm_bound_integral_component: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" - fixes g :: "'n \ 'b::ordered_euclidean_space" + fixes f :: "'n::euclidean_space \ 'a::banach" + fixes g :: "'n \ 'b::euclidean_space" assumes "f integrable_on s" and "g integrable_on s" and "\x\s. norm(f x) \ (g x)\k" @@ -11037,8 +11207,8 @@ qed lemma has_integral_norm_bound_integral_component: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" - fixes g :: "'n \ 'b::ordered_euclidean_space" + fixes f :: "'n::euclidean_space \ 'a::banach" + fixes g :: "'n \ 'b::euclidean_space" assumes "(f has_integral i) s" and "(g has_integral j) s" and "\x\s. norm (f x) \ (g x)\k" @@ -11049,7 +11219,7 @@ by auto lemma absolutely_integrable_le: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f absolutely_integrable_on s" shows "norm (integral s f) \ integral s (\x. norm (f x))" apply (rule integral_norm_bound_integral) @@ -11092,14 +11262,14 @@ done lemma absolutely_integrable_on_subinterval: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" shows "f absolutely_integrable_on s \ - {a..b} \ s \ f absolutely_integrable_on {a..b}" + cbox a b \ s \ f absolutely_integrable_on cbox a b" unfolding absolutely_integrable_on_def - by (metis integrable_on_subinterval) + by (metis integrable_on_subcbox) lemma absolutely_integrable_bounded_variation: - fixes f :: "'n::ordered_euclidean_space \ 'a::banach" + fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f absolutely_integrable_on UNIV" obtains B where "\d. d division_of (\d) \ setsum (\k. norm(integral k f)) d \ B" apply (rule that[of "integral UNIV (\x. norm (f x))"]) @@ -11140,12 +11310,12 @@ done lemma bounded_variation_absolutely_integrable_interval: - fixes f :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" - assumes "f integrable_on {a..b}" - and *: "\d. d division_of {a..b} \ setsum (\k. norm(integral k f)) d \ B" - shows "f absolutely_integrable_on {a..b}" -proof - - let ?f = "\d. \k\d. norm (integral k f)" and ?D = "{d. d division_of {a..b}}" + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + assumes "f integrable_on cbox a b" + and *: "\d. d division_of (cbox a b) \ setsum (\k. norm(integral k f)) d \ B" + shows "f absolutely_integrable_on cbox a b" +proof - + let ?f = "\d. \k\d. norm (integral k f)" and ?D = "{d. d division_of (cbox a b)}" have D_1: "?D \ {}" by (rule elementary_interval[of a b]) auto have D_2: "bdd_above (?f`?D)" @@ -11160,7 +11330,7 @@ proof safe case goal1 then have "?S - e / 2 < ?S" by simp - then obtain d where d: "d division_of {a..b}" "?S - e / 2 < (\k\d. norm (integral k f))" + then obtain d where d: "d division_of (cbox a b)" "?S - e / 2 < (\k\d. norm (integral k f))" unfolding less_cSUP_iff[OF D] by auto note d' = division_ofD[OF this(1)] @@ -11199,7 +11369,7 @@ using k(1) by auto fix p - assume "p tagged_division_of {a..b}" and "?g fine p" + assume "p tagged_division_of (cbox a b)" and "?g fine p" note p = this(1) conjunctD2[OF this(2)[unfolded fine_inter]] note p' = tagged_division_ofD[OF p(1)] def p' \ "{(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" @@ -11207,7 +11377,7 @@ using p(2) unfolding p'_def fine_def by auto - have p'': "p' tagged_division_of {a..b}" + have p'': "p' tagged_division_of (cbox a b)" apply (rule tagged_division_ofI) proof - show "finite p'" @@ -11226,9 +11396,9 @@ then have "\i l. x \ i \ i \ d \ (x, l) \ p \ k = i \ l" unfolding p'_def by auto then guess i l by (elim exE) note il=conjunctD4[OF this] - show "x \ k" and "k \ {a..b}" + show "x \ k" and "k \ cbox a b" using p'(2-3)[OF il(3)] il by auto - show "\a b. k = {a..b}" + show "\a b. k = cbox a b" unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] apply safe unfolding inter_interval @@ -11253,9 +11423,9 @@ then show "interior k1 \ interior k2 = {}" unfolding il1 il2 by auto next - have *: "\(x, X) \ p'. X \ {a..b}" + have *: "\(x, X) \ p'. X \ cbox a b" unfolding p'_def using d' by auto - show "\{k. \x. (x, k) \ p'} = {a..b}" + show "\{k. \x. (x, k) \ p'} = cbox a b" apply rule apply (rule Union_least) unfolding mem_Collect_eq @@ -11264,7 +11434,7 @@ apply safe proof - fix y - assume y: "y \ {a..b}" + assume y: "y \ cbox a b" then have "\x l. (x, l) \ p \ y\l" unfolding p'(6)[symmetric] by auto then guess x l by (elim exE) note xl=conjunctD2[OF this] @@ -11367,9 +11537,9 @@ case goal1 note k=this from d'(4)[OF this] guess u v by (elim exE) note uv=this - def d' \ "{{u..v} \ l |l. l \ snd ` p \ {u..v} \ l \ {}}" + def d' \ "{cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" note uvab = d'(2)[OF k[unfolded uv]] - have "d' division_of {u..v}" + have "d' division_of cbox u v" apply (subst d'_def) apply (rule division_inter_1) apply (rule division_of_tagged_division[OF p(1)]) @@ -11378,7 +11548,7 @@ then have "norm (integral k f) \ setsum (\k. norm (integral k f)) d'" unfolding uv apply (subst integral_combine_division_topdown[of _ _ d']) - apply (rule integrable_on_subinterval[OF assms(1) uvab]) + apply (rule integrable_on_subcbox[OF assms(1) uvab]) apply assumption apply (rule setsum_norm_le) apply auto @@ -11392,10 +11562,10 @@ apply blast proof case goal1 - then have "i \ {{u..v} \ l |l. l \ snd ` p}" + then have "i \ {cbox u v \ l |l. l \ snd ` p}" by auto from this[unfolded mem_Collect_eq] guess l .. note l=this - then have "{u..v} \ l = {}" + then have "cbox u v \ l = {}" using goal1 by auto then show ?case using l by auto @@ -11578,7 +11748,7 @@ assume as: "(x, l) \ p" note xl = p'(2-4)[OF this] from this(3) guess u v by (elim exE) note uv=this - have "(\i\d. \content (l \ i)\) = (\k\d. content (k \ {u..v}))" + have "(\i\d. \content (l \ i)\) = (\k\d. content (k \ cbox u v))" apply (rule setsum_cong2) apply (drule d'(4)) apply safe @@ -11587,7 +11757,7 @@ apply (subst abs_of_nonneg) apply auto done - also have "\ = setsum content {k \ {u..v}| k. k \ d}" + also have "\ = setsum content {k \ cbox u v| k. k \ d}" unfolding simple_image apply (rule setsum_reindex_nonzero[unfolded o_def,symmetric]) apply (rule d') @@ -11595,19 +11765,19 @@ case goal1 from d'(4)[OF this(1)] d'(4)[OF this(2)] guess u1 v1 u2 v2 by (elim exE) note uv=this - have "{} = interior ((k \ y) \ {u..v})" + have "{} = interior ((k \ y) \ cbox u v)" apply (subst interior_inter) using d'(5)[OF goal1(1-3)] apply auto done - also have "\ = interior (y \ (k \ {u..v}))" + also have "\ = interior (y \ (k \ cbox u v))" by auto - also have "\ = interior (k \ {u..v})" + also have "\ = interior (k \ cbox u v)" unfolding goal1(4) by auto finally show ?case unfolding uv inter_interval content_eq_0_interior .. qed - also have "\ = setsum content {{u..v} \ k |k. k \ d \ {u..v} \ k \ {}}" + also have "\ = setsum content {cbox u v \ k |k. k \ d \ cbox u v \ k \ {}}" apply (rule setsum_mono_zero_right) unfolding simple_image apply (rule finite_imageI) @@ -11618,13 +11788,13 @@ proof - case goal1 from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this - have "interior (k \ {u..v}) \ {}" + have "interior (k \ cbox u v) \ {}" using goal1(2) unfolding ab inter_interval content_eq_0_interior by auto then show ?case using goal1(1) - using interior_subset[of "k \ {u..v}"] + using interior_subset[of "k \ cbox u v"] by auto qed finally show "(\i\d. \content (l \ i)\ * norm (f x)) = content l *\<^sub>R norm (f x)" @@ -11642,7 +11812,7 @@ qed lemma bounded_variation_absolutely_integrable: - fixes f :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on UNIV" and "\d. d division_of (\d) \ setsum (\k. norm (integral k f)) d \ B" shows "f absolutely_integrable_on UNIV" @@ -11654,9 +11824,9 @@ by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp note D = D_1 D_2 let ?S = "SUP d:?D. ?f d" - have f_int: "\a b. f absolutely_integrable_on {a..b}" + have f_int: "\a b. f absolutely_integrable_on cbox a b" apply (rule bounded_variation_absolutely_integrable_interval[where B=B]) - apply (rule integrable_on_subinterval[OF assms(1)]) + apply (rule integrable_on_subcbox[OF assms(1)]) defer apply safe apply (rule assms(2)[rule_format]) @@ -11696,10 +11866,10 @@ apply safe proof - fix a b :: 'n - assume ab: "ball 0 (K + 1) \ {a..b}" + assume ab: "ball 0 (K + 1) \ cbox a b" have *: "\s s1. ?S - e < s1 \ s1 \ s \ s < ?S + e \ abs (s - ?S) < e" by arith - show "norm (integral {a..b} (\x. if x \ UNIV then norm (f x) else 0) - ?S) < e" + show "norm (integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) - ?S) < e" unfolding real_norm_def apply (rule *[rule_format]) apply safe @@ -11720,10 +11890,10 @@ using f_int apply auto done - also have "\ \ integral {a..b} (\x. if x \ UNIV then norm (f x) else 0)" + also have "\ \ integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0)" proof - case goal1 - have "\d \ {a..b}" + have "\d \ cbox a b" apply rule apply (drule K(2)[rule_format]) apply (rule ab[unfolded subset_eq,rule_format]) @@ -11735,7 +11905,7 @@ apply rule apply (rule integral_subset_le) defer - apply (rule integrable_on_subdivision[of _ _ _ "{a..b}"]) + apply (rule integrable_on_subdivision[of _ _ _ "cbox a b"]) apply (rule d) using f_int[of a b] apply auto @@ -11748,21 +11918,21 @@ have "e/2>0" using `e > 0` by auto from * [OF this] obtain d1 where - d1: "gauge d1" "\p. p tagged_division_of {a..b} \ d1 fine p \ - norm ((\(x, k)\p. content k *\<^sub>R norm (f x)) - integral {a..b} (\x. norm (f x))) < e / 2" + d1: "gauge d1" "\p. p tagged_division_of (cbox a b) \ d1 fine p \ + norm ((\(x, k)\p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm (f x))) < e / 2" by auto from henstock_lemma [OF f(1) `e/2>0`] obtain d2 where - d2: "gauge d2" "\p. p tagged_partial_division_of {a..b} \ d2 fine p \ + d2: "gauge d2" "\p. p tagged_partial_division_of (cbox a b) \ d2 fine p \ (\(x, k)\p. norm (content k *\<^sub>R f x - integral k f)) < e / 2" by blast obtain p where - p: "p tagged_division_of {a..b}" "d1 fine p" "d2 fine p" + p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p" by (rule fine_division_exists [OF gauge_inter [OF d1(1) d2(1)], of a b]) (auto simp add: fine_inter) have *: "\sf sf' si di. sf' = sf \ si \ ?S \ abs (sf - si) < e / 2 \ abs (sf' - di) < e / 2 \ di < ?S + e" by arith - show "integral {a..b} (\x. if x \ UNIV then norm (f x) else 0) < ?S + e" + show "integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) < ?S + e" apply (subst if_P) apply rule proof (rule *[rule_format]) @@ -11774,7 +11944,7 @@ unfolding tagged_division_of_def split_def apply auto done - show "abs ((\(x, k)\p. content k *\<^sub>R norm (f x)) - integral {a..b} (\x. norm(f x))) < e / 2" + show "abs ((\(x, k)\p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm(f x))) < e / 2" using d1(2)[rule_format,OF conjI[OF p(1,2)]] by (simp only: real_norm_def) show "(\(x, k)\p. content k *\<^sub>R norm (f x)) = (\(x, k)\p. norm (content k *\<^sub>R f x))" @@ -11786,7 +11956,7 @@ apply auto done show "(\(x, k)\p. norm (integral k f)) \ ?S" - using partial_division_of_tagged_division[of p "{a..b}"] p(1) + using partial_division_of_tagged_division[of p "cbox a b"] p(1) apply (subst setsum_over_tagged_division_lemma[OF p(1)]) apply (simp add: integral_null) apply (intro cSUP_upper2[OF D(2), of "snd ` p"]) @@ -11804,7 +11974,7 @@ unfolding absolutely_integrable_on_def if_distrib norm_zero integrable_restrict_univ .. lemma absolutely_integrable_add[intro]: - fixes f g :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f g :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on s" and "g absolutely_integrable_on s" shows "(\x. f x + g x) absolutely_integrable_on s" @@ -11835,7 +12005,7 @@ have "\k. k \ d \ f integrable_on k \ g integrable_on k" apply (drule division_ofD(4)[OF goal1]) apply safe - apply (rule_tac[!] integrable_on_subinterval[of _ UNIV]) + apply (rule_tac[!] integrable_on_subcbox[of _ UNIV]) using assms apply auto done @@ -11856,7 +12026,7 @@ qed lemma absolutely_integrable_sub[intro]: - fixes f g :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f g :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on s" and "g absolutely_integrable_on s" shows "(\x. f x - g x) absolutely_integrable_on s" @@ -11864,8 +12034,8 @@ by (simp add: algebra_simps) lemma absolutely_integrable_linear: - fixes f :: "'m::ordered_euclidean_space \ 'n::ordered_euclidean_space" - and h :: "'n::ordered_euclidean_space \ 'p::ordered_euclidean_space" + fixes f :: "'m::euclidean_space \ 'n::euclidean_space" + and h :: "'n::euclidean_space \ 'p::euclidean_space" assumes "f absolutely_integrable_on s" and "bounded_linear h" shows "(h \ f) absolutely_integrable_on s" @@ -11899,7 +12069,7 @@ guess u v by (elim exE) note uv=this have *: "f integrable_on k" unfolding uv - apply (rule integrable_on_subinterval[of _ UNIV]) + apply (rule integrable_on_subcbox[of _ UNIV]) using assms apply auto done @@ -11919,7 +12089,7 @@ qed lemma absolutely_integrable_setsum: - fixes f :: "'a \ 'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f :: "'a \ 'n::euclidean_space \ 'm::euclidean_space" assumes "finite t" and "\a. a \ t \ (f a) absolutely_integrable_on s" shows "(\x. setsum (\a. f a x) t) absolutely_integrable_on s" @@ -11946,8 +12116,8 @@ qed lemma absolutely_integrable_vector_abs: - fixes f :: "'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - and T :: "'c::ordered_euclidean_space \ 'b" + fixes f :: "'a::euclidean_space => 'b::euclidean_space" + and T :: "'c::euclidean_space \ 'b" assumes f: "f absolutely_integrable_on s" shows "(\x. (\i\Basis. abs(f x\T i) *\<^sub>R i)) absolutely_integrable_on s" (is "?Tf absolutely_integrable_on s") @@ -11969,7 +12139,7 @@ qed lemma absolutely_integrable_max: - fixes f g :: "'m::ordered_euclidean_space \ 'n::ordered_euclidean_space" + fixes f g :: "'m::euclidean_space \ 'n::euclidean_space" assumes "f absolutely_integrable_on s" and "g absolutely_integrable_on s" shows "(\x. (\i\Basis. max (f(x)\i) (g(x)\i) *\<^sub>R i)::'n) absolutely_integrable_on s" @@ -11985,7 +12155,7 @@ qed lemma absolutely_integrable_min: - fixes f g::"'m::ordered_euclidean_space \ 'n::ordered_euclidean_space" + fixes f g::"'m::euclidean_space \ 'n::euclidean_space" assumes "f absolutely_integrable_on s" and "g absolutely_integrable_on s" shows "(\x. (\i\Basis. min (f(x)\i) (g(x)\i) *\<^sub>R i)::'n) absolutely_integrable_on s" @@ -12002,7 +12172,7 @@ qed lemma absolutely_integrable_abs_eq: - fixes f::"'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f::"'n::euclidean_space \ 'm::euclidean_space" shows "f absolutely_integrable_on s \ f integrable_on s \ (\x. (\i\Basis. abs(f x\i) *\<^sub>R i)::'m) integrable_on s" (is "?l = ?r") @@ -12061,8 +12231,8 @@ defer apply (rule_tac[1-2] integral_component_le[OF i]) apply (rule integrable_neg) - using integrable_on_subinterval[OF assms(1),of a b] - integrable_on_subinterval[OF assms(2),of a b] i + using integrable_on_subcbox[OF assms(1),of a b] + integrable_on_subcbox[OF assms(2),of a b] i unfolding ab apply auto done @@ -12089,7 +12259,7 @@ qed lemma nonnegative_absolutely_integrable: - fixes f :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "\x\s. \i\Basis. 0 \ f x \ i" and "f integrable_on s" shows "f absolutely_integrable_on s" @@ -12102,7 +12272,7 @@ done lemma absolutely_integrable_integrable_bound: - fixes f :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "\x\s. norm (f x) \ g x" and "f integrable_on s" and "g integrable_on s" @@ -12132,7 +12302,7 @@ apply (rule integral_norm_bound_integral) apply (drule_tac[!] d'(4)) apply safe - apply (rule_tac[1-2] integrable_on_subinterval) + apply (rule_tac[1-2] integrable_on_subcbox) using assms apply auto done @@ -12142,7 +12312,7 @@ apply safe apply (drule d'(4)) apply safe - apply (rule integrable_on_subinterval[OF assms(3)]) + apply (rule integrable_on_subcbox[OF assms(3)]) apply auto done also have "\ \ integral UNIV g" @@ -12160,7 +12330,7 @@ qed lemma absolutely_integrable_integrable_bound_real: - fixes f :: "'n::ordered_euclidean_space \ real" + fixes f :: "'n::euclidean_space \ real" assumes "\x\s. norm (f x) \ g x" and "f integrable_on s" and "g integrable_on s" @@ -12172,8 +12342,8 @@ done lemma absolutely_integrable_absolutely_integrable_bound: - fixes f :: "'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" - and g :: "'n::ordered_euclidean_space \ 'p::ordered_euclidean_space" + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + and g :: "'n::euclidean_space \ 'p::euclidean_space" assumes "\x\s. norm (f x) \ norm (g x)" and "f integrable_on s" and "g absolutely_integrable_on s" @@ -12264,7 +12434,7 @@ subsection {* Dominated convergence *} lemma dominated_convergence: - fixes f :: "nat \ 'n::ordered_euclidean_space \ real" + fixes f :: "nat \ 'n::euclidean_space \ real" assumes "\k. (f k) integrable_on s" "h integrable_on s" and "\k. \x \ s. norm (f k x) \ h x" and "\x \ s. ((\k. f k x) ---> g x) sequentially" diff -r 2666cd7d380c -r 0268784f60da src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy Mon Mar 17 21:56:32 2014 +0100 +++ b/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy Tue Mar 18 10:12:57 2014 +0100 @@ -175,31 +175,31 @@ eucl_le[where 'a='b] abs_prod_def abs_inner) -subsection {* Intervals *} +subsection {* Boxes *} -lemma interval: - fixes a :: "'a::ordered_euclidean_space" +lemma box: + fixes a :: "'a::euclidean_space" shows "box a b = {x::'a. \i\Basis. a\i < x\i \ x\i < b\i}" - and "{a .. b} = {x::'a. \i\Basis. a\i \ x\i \ x\i \ b\i}" - by (auto simp add:set_eq_iff eucl_le[where 'a='a] box_def) + and "cbox a b = {x::'a. \i\Basis. a\i \ x\i \ x\i \ b\i}" + by (auto simp add:set_eq_iff box_def cbox_def) -lemma mem_interval: - fixes a :: "'a::ordered_euclidean_space" +lemma mem_box: + fixes a :: "'a::euclidean_space" shows "x \ box a b \ (\i\Basis. a\i < x\i \ x\i < b\i)" - and "x \ {a .. b} \ (\i\Basis. a\i \ x\i \ x\i \ b\i)" - using interval[of a b] + and "x \ cbox a b \ (\i\Basis. a\i \ x\i \ x\i \ b\i)" + using box[of a b] by auto -lemma interval_eq_empty: - fixes a :: "'a::ordered_euclidean_space" +lemma box_eq_empty: + fixes a :: "'a::euclidean_space" shows "(box a b = {} \ (\i\Basis. b\i \ a\i))" (is ?th1) - and "({a .. b} = {} \ (\i\Basis. b\i < a\i))" (is ?th2) + and "(cbox a b = {} \ (\i\Basis. b\i < a\i))" (is ?th2) proof - { fix i x assume i: "i\Basis" and as:"b\i \ a\i" and x:"x\box a b" then have "a \ i < x \ i \ x \ i < b \ i" - unfolding mem_interval by auto + unfolding mem_box by (auto simp: box_def) then have "a\i < b\i" by auto then have False using as by auto } @@ -216,15 +216,15 @@ by (auto simp: inner_add_left) } then have "box a b \ {}" - using mem_interval(1)[of "?x" a b] by auto + using mem_box(1)[of "?x" a b] by auto } ultimately show ?th1 by blast { fix i x - assume i: "i \ Basis" and as:"b\i < a\i" and x:"x\{a .. b}" + assume i: "i \ Basis" and as:"b\i < a\i" and x:"x\cbox a b" then have "a \ i \ x \ i \ x \ i \ b \ i" - unfolding mem_interval by auto + unfolding mem_box by auto then have "a\i \ b\i" by auto then have False using as by auto } @@ -240,57 +240,58 @@ then have "a\i \ ((1/2) *\<^sub>R (a+b)) \ i" "((1/2) *\<^sub>R (a+b)) \ i \ b\i" by (auto simp: inner_add_left) } - then have "{a .. b} \ {}" - using mem_interval(2)[of "?x" a b] by auto + then have "cbox a b \ {}" + using mem_box(2)[of "?x" a b] by auto } ultimately show ?th2 by blast qed -lemma interval_ne_empty: - fixes a :: "'a::ordered_euclidean_space" - shows "{a .. b} \ {} \ (\i\Basis. a\i \ b\i)" +lemma box_ne_empty: + fixes a :: "'a::euclidean_space" + shows "cbox a b \ {} \ (\i\Basis. a\i \ b\i)" and "box a b \ {} \ (\i\Basis. a\i < b\i)" - unfolding interval_eq_empty[of a b] by fastforce+ + unfolding box_eq_empty[of a b] by fastforce+ -lemma interval_sing: - fixes a :: "'a::ordered_euclidean_space" - shows "{a .. a} = {a}" - and "box a a = {}" - unfolding set_eq_iff mem_interval eq_iff [symmetric] - by (auto intro: euclidean_eqI simp: ex_in_conv) +lemma + fixes a :: "'a::euclidean_space" + shows cbox_sing: "cbox a a = {a}" + and box_sing: "box a a = {}" + unfolding set_eq_iff mem_box eq_iff [symmetric] + by (auto intro!: euclidean_eqI[where 'a='a]) + (metis all_not_in_conv nonempty_Basis) -lemma subset_interval_imp: - fixes a :: "'a::ordered_euclidean_space" - shows "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ {c .. d} \ {a .. b}" - and "(\i\Basis. a\i < c\i \ d\i < b\i) \ {c .. d} \ box a b" - and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ {a .. b}" - and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ box a b" - unfolding subset_eq[unfolded Ball_def] unfolding mem_interval - by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+ +lemma subset_box_imp: + fixes a :: "'a::euclidean_space" + shows "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ cbox c d \ cbox a b" + and "(\i\Basis. a\i < c\i \ d\i < b\i) \ cbox c d \ box a b" + and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ cbox a b" + and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ box a b" + unfolding subset_eq[unfolded Ball_def] unfolding mem_box + by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+ -lemma interval_open_subset_closed: - fixes a :: "'a::ordered_euclidean_space" - shows "box a b \ {a .. b}" - unfolding subset_eq [unfolded Ball_def] mem_interval +lemma box_subset_cbox: + fixes a :: "'a::euclidean_space" + shows "box a b \ cbox a b" + unfolding subset_eq [unfolded Ball_def] mem_box by (fast intro: less_imp_le) -lemma subset_interval: - fixes a :: "'a::ordered_euclidean_space" - shows "{c .. d} \ {a .. b} \ (\i\Basis. c\i \ d\i) --> (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th1) - and "{c .. d} \ box a b \ (\i\Basis. c\i \ d\i) --> (\i\Basis. a\i < c\i \ d\i < b\i)" (is ?th2) - and "box c d \ {a .. b} \ (\i\Basis. c\i < d\i) --> (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th3) +lemma subset_box: + fixes a :: "'a::euclidean_space" + shows "cbox c d \ cbox a b \ (\i\Basis. c\i \ d\i) --> (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th1) + and "cbox c d \ box a b \ (\i\Basis. c\i \ d\i) --> (\i\Basis. a\i < c\i \ d\i < b\i)" (is ?th2) + and "box c d \ cbox a b \ (\i\Basis. c\i < d\i) --> (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th3) and "box c d \ box a b \ (\i\Basis. c\i < d\i) --> (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th4) proof - show ?th1 - unfolding subset_eq and Ball_def and mem_interval + unfolding subset_eq and Ball_def and mem_box by (auto intro: order_trans) show ?th2 - unfolding subset_eq and Ball_def and mem_interval + unfolding subset_eq and Ball_def and mem_box by (auto intro: le_less_trans less_le_trans order_trans less_imp_le) { - assume as: "box c d \ {a .. b}" "\i\Basis. c\i < d\i" + assume as: "box c d \ cbox a b" "\i\Basis. c\i < d\i" then have "box c d \ {}" - unfolding interval_eq_empty by auto + unfolding box_eq_empty by auto fix i :: 'a assume i: "i \ Basis" (** TODO combine the following two parts as done in the HOL_light version. **) @@ -307,10 +308,10 @@ done } then have "?x\box c d" - using i unfolding mem_interval by auto + using i unfolding mem_box by auto moreover - have "?x \ {a .. b}" - unfolding mem_interval + have "?x \ cbox a b" + unfolding mem_box apply auto apply (rule_tac x=i in bexI) using as(2)[THEN bspec[where x=i]] and as2 i @@ -333,10 +334,10 @@ done } then have "?x\box c d" - unfolding mem_interval by auto + unfolding mem_box by auto moreover - have "?x\{a .. b}" - unfolding mem_interval + have "?x\cbox a b" + unfolding mem_box apply auto apply (rule_tac x=i in bexI) using as(2)[THEN bspec[where x=i]] and as2 using i @@ -349,10 +350,10 @@ have "a\i \ c\i \ d\i \ b\i" by auto } note part1 = this show ?th3 - unfolding subset_eq and Ball_def and mem_interval + unfolding subset_eq and Ball_def and mem_box apply (rule, rule, rule, rule) apply (rule part1) - unfolding subset_eq and Ball_def and mem_interval + unfolding subset_eq and Ball_def and mem_box prefer 4 apply auto apply (erule_tac x=xa in allE, erule_tac x=xa in allE, fastforce)+ @@ -361,16 +362,16 @@ assume as: "box c d \ box a b" "\i\Basis. c\i < d\i" fix i :: 'a assume i:"i\Basis" - from as(1) have "box c d \ {a..b}" - using interval_open_subset_closed[of a b] by auto + from as(1) have "box c d \ cbox a b" + using box_subset_cbox[of a b] by auto then have "a\i \ c\i \ d\i \ b\i" using part1 and as(2) using i by auto } note * = this show ?th4 - unfolding subset_eq and Ball_def and mem_interval + unfolding subset_eq and Ball_def and mem_box apply (rule, rule, rule, rule) apply (rule *) - unfolding subset_eq and Ball_def and mem_interval + unfolding subset_eq and Ball_def and mem_box prefer 4 apply auto apply (erule_tac x=xa in allE, simp)+ @@ -378,68 +379,68 @@ qed lemma inter_interval: - fixes a :: "'a::ordered_euclidean_space" - shows "{a .. b} \ {c .. d} = - {(\i\Basis. max (a\i) (c\i) *\<^sub>R i) .. (\i\Basis. min (b\i) (d\i) *\<^sub>R i)}" - unfolding set_eq_iff and Int_iff and mem_interval + fixes a :: "'a::euclidean_space" + shows "cbox a b \ cbox c d = + cbox (\i\Basis. max (a\i) (c\i) *\<^sub>R i) (\i\Basis. min (b\i) (d\i) *\<^sub>R i)" + unfolding set_eq_iff and Int_iff and mem_box by auto lemma disjoint_interval: - fixes a::"'a::ordered_euclidean_space" - shows "{a .. b} \ {c .. d} = {} \ (\i\Basis. (b\i < a\i \ d\i < c\i \ b\i < c\i \ d\i < a\i))" (is ?th1) - and "{a .. b} \ box c d = {} \ (\i\Basis. (b\i < a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th2) - and "box a b \ {c .. d} = {} \ (\i\Basis. (b\i \ a\i \ d\i < c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th3) + fixes a::"'a::euclidean_space" + shows "cbox a b \ cbox c d = {} \ (\i\Basis. (b\i < a\i \ d\i < c\i \ b\i < c\i \ d\i < a\i))" (is ?th1) + and "cbox a b \ box c d = {} \ (\i\Basis. (b\i < a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th2) + and "box a b \ cbox c d = {} \ (\i\Basis. (b\i \ a\i \ d\i < c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th3) and "box a b \ box c d = {} \ (\i\Basis. (b\i \ a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th4) proof - let ?z = "(\i\Basis. (((max (a\i) (c\i)) + (min (b\i) (d\i))) / 2) *\<^sub>R i)::'a" have **: "\P Q. (\i :: 'a. i \ Basis \ Q ?z i \ P i) \ (\i x :: 'a. i \ Basis \ P i \ Q x i) \ (\x. \i\Basis. Q x i) \ (\i\Basis. P i)" by blast - note * = set_eq_iff Int_iff empty_iff mem_interval ball_conj_distrib[symmetric] eq_False ball_simps(10) + note * = set_eq_iff Int_iff empty_iff mem_box ball_conj_distrib[symmetric] eq_False ball_simps(10) show ?th1 unfolding * by (intro **) auto show ?th2 unfolding * by (intro **) auto show ?th3 unfolding * by (intro **) auto show ?th4 unfolding * by (intro **) auto qed -(* Moved interval_open_subset_closed a bit upwards *) +(* Moved box_subset_cbox a bit upwards *) -lemma open_interval[intro]: - fixes a b :: "'a::ordered_euclidean_space" +lemma open_box[intro]: + fixes a b :: "'a::euclidean_space" shows "open (box a b)" proof - have "open (\i\Basis. (\x. x\i) -` {a\i<..i})" by (intro open_INT finite_lessThan ballI continuous_open_vimage allI linear_continuous_at open_real_greaterThanLessThan finite_Basis bounded_linear_inner_left) also have "(\i\Basis. (\x. x\i) -` {a\i<..i}) = box a b" - by (auto simp add: interval) + by (auto simp add: box) finally show "open (box a b)" . qed -lemma closed_interval[intro]: - fixes a b :: "'a::ordered_euclidean_space" - shows "closed {a .. b}" +lemma closed_cbox[intro]: + fixes a b :: "'a::euclidean_space" + shows "closed (cbox a b)" proof - have "closed (\i\Basis. (\x. x\i) -` {a\i .. b\i})" by (intro closed_INT ballI continuous_closed_vimage allI linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left) - also have "(\i\Basis. (\x. x\i) -` {a\i .. b\i}) = {a .. b}" - by (auto simp add: eucl_le [where 'a='a]) - finally show "closed {a .. b}" . + also have "(\i\Basis. (\x. x\i) -` {a\i .. b\i}) = cbox a b" + by (auto simp add: cbox_def) + finally show "closed (cbox a b)" . qed -lemma interior_closed_interval [intro]: - fixes a b :: "'a::ordered_euclidean_space" - shows "interior {a..b} = box a b" (is "?L = ?R") +lemma interior_cbox [intro]: + fixes a b :: "'a::euclidean_space" + shows "interior (cbox a b) = box a b" (is "?L = ?R") proof(rule subset_antisym) show "?R \ ?L" - using interval_open_subset_closed open_interval + using box_subset_cbox open_box by (rule interior_maximal) { fix x - assume "x \ interior {a..b}" - then obtain s where s: "open s" "x \ s" "s \ {a..b}" .. - then obtain e where "e>0" and e:"\x'. dist x' x < e \ x' \ {a..b}" + assume "x \ interior (cbox a b)" + then obtain s where s: "open s" "x \ s" "s \ cbox a b" .. + then obtain e where "e>0" and e:"\x'. dist x' x < e \ x' \ cbox a b" unfolding open_dist and subset_eq by auto { fix i :: 'a @@ -455,7 +456,7 @@ then have "a \ i \ (x - (e / 2) *\<^sub>R i) \ i" and "(x + (e / 2) *\<^sub>R i) \ i \ b \ i" using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] - unfolding mem_interval + unfolding mem_box using i by blast+ then have "a \ i < x \ i" and "x \ i < b \ i" @@ -463,14 +464,14 @@ by (auto simp: inner_diff_left inner_Basis inner_add_left) } then have "x \ box a b" - unfolding mem_interval by auto + unfolding mem_box by auto } then show "?L \ ?R" .. qed -lemma bounded_closed_interval: - fixes a :: "'a::ordered_euclidean_space" - shows "bounded {a .. b}" +lemma bounded_cbox: + fixes a :: "'a::euclidean_space" + shows "bounded (cbox a b)" proof - let ?b = "\i\Basis. \a\i\ + \b\i\" { @@ -491,30 +492,30 @@ using norm_le_l1[of x] by auto } then show ?thesis - unfolding interval and bounded_iff by auto + unfolding box and bounded_iff by auto qed lemma bounded_interval: - fixes a :: "'a::ordered_euclidean_space" - shows "bounded {a .. b} \ bounded (box a b)" - using bounded_closed_interval[of a b] - using interval_open_subset_closed[of a b] - using bounded_subset[of "{a..b}" "box a b"] + fixes a :: "'a::euclidean_space" + shows "bounded (cbox a b) \ bounded (box a b)" + using bounded_cbox[of a b] + using box_subset_cbox[of a b] + using bounded_subset[of "cbox a b" "box a b"] by simp lemma not_interval_univ: - fixes a :: "'a::ordered_euclidean_space" - shows "{a .. b} \ UNIV \ box a b \ UNIV" + fixes a :: "'a::euclidean_space" + shows "cbox a b \ UNIV \ box a b \ UNIV" using bounded_interval[of a b] by auto -lemma compact_interval: - fixes a :: "'a::ordered_euclidean_space" - shows "compact {a .. b}" - using bounded_closed_imp_seq_compact[of "{a..b}"] using bounded_interval[of a b] +lemma compact_cbox: + fixes a :: "'a::euclidean_space" + shows "compact (cbox a b)" + using bounded_closed_imp_seq_compact[of "cbox a b"] using bounded_interval[of a b] by (auto simp: compact_eq_seq_compact_metric) lemma open_interval_midpoint: - fixes a :: "'a::ordered_euclidean_space" + fixes a :: "'a::euclidean_space" assumes "box a b \ {}" shows "((1/2) *\<^sub>R (a + b)) \ box a b" proof - @@ -522,15 +523,15 @@ fix i :: 'a assume "i \ Basis" then have "a \ i < ((1 / 2) *\<^sub>R (a + b)) \ i \ ((1 / 2) *\<^sub>R (a + b)) \ i < b \ i" - using assms[unfolded interval_ne_empty, THEN bspec[where x=i]] by (auto simp: inner_add_left) + using assms[unfolded box_ne_empty, THEN bspec[where x=i]] by (auto simp: inner_add_left) } - then show ?thesis unfolding mem_interval by auto + then show ?thesis unfolding mem_box by auto qed lemma open_closed_interval_convex: - fixes x :: "'a::ordered_euclidean_space" + fixes x :: "'a::euclidean_space" assumes x: "x \ box a b" - and y: "y \ {a .. b}" + and y: "y \ cbox a b" and e: "0 < e" "e \ 1" shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ box a b" proof - @@ -543,9 +544,9 @@ apply (rule add_less_le_mono) using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all - using x unfolding mem_interval using i + using x unfolding mem_box using i apply simp - using y unfolding mem_interval using i + using y unfolding mem_box using i apply simp done finally have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i" @@ -559,11 +560,11 @@ using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all using x - unfolding mem_interval + unfolding mem_box using i apply simp using y - unfolding mem_interval + unfolding mem_box using i apply simp done @@ -574,23 +575,23 @@ by auto } then show ?thesis - unfolding mem_interval by auto + unfolding mem_box by auto qed notation eucl_less (infix " {}" - shows "closure (box a b) = {a .. b}" + fixes a :: "'a::euclidean_space" + assumes "box a b \ {}" + shows "closure (box a b) = cbox a b" proof - have ab: "a {a .. b}" + assume as:"x \ cbox a b" def f \ "\n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)" { fix n @@ -602,7 +603,7 @@ by (auto simp add: algebra_simps) then have "f n a. s \ box (-a) a" proof - @@ -665,38 +666,38 @@ by auto } then show ?thesis - by (auto intro: exI[where x=a] simp add: interval) + by (auto intro: exI[where x=a] simp add: box) qed lemma bounded_subset_open_interval: - fixes s :: "('a::ordered_euclidean_space) set" + fixes s :: "('a::euclidean_space) set" shows "bounded s \ (\a b. s \ box a b)" by (auto dest!: bounded_subset_open_interval_symmetric) -lemma bounded_subset_closed_interval_symmetric: - fixes s :: "('a::ordered_euclidean_space) set" - assumes "bounded s" - shows "\a. s \ {-a .. a}" +lemma bounded_subset_cbox_symmetric: + fixes s :: "('a::euclidean_space) set" + assumes "bounded s" + shows "\a. s \ cbox (-a) a" proof - obtain a where "s \ box (-a) a" using bounded_subset_open_interval_symmetric[OF assms] by auto then show ?thesis - using interval_open_subset_closed[of "-a" a] by auto + using box_subset_cbox[of "-a" a] by auto qed -lemma bounded_subset_closed_interval: - fixes s :: "('a::ordered_euclidean_space) set" - shows "bounded s \ \a b. s \ {a .. b}" - using bounded_subset_closed_interval_symmetric[of s] by auto +lemma bounded_subset_cbox: + fixes s :: "('a::euclidean_space) set" + shows "bounded s \ \a b. s \ cbox a b" + using bounded_subset_cbox_symmetric[of s] by auto lemma frontier_closed_interval: - fixes a b :: "'a::ordered_euclidean_space" - shows "frontier {a .. b} = {a .. b} - box a b" - unfolding frontier_def unfolding interior_closed_interval and closure_closed[OF closed_interval] .. + fixes a b :: "'a::euclidean_space" + shows "frontier (cbox a b) = cbox a b - box a b" + unfolding frontier_def unfolding interior_cbox and closure_closed[OF closed_cbox] .. lemma frontier_open_interval: - fixes a b :: "'a::ordered_euclidean_space" - shows "frontier (box a b) = (if box a b = {} then {} else {a .. b} - box a b)" + fixes a b :: "'a::euclidean_space" + shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)" proof (cases "box a b = {}") case True then show ?thesis @@ -704,16 +705,16 @@ next case False then show ?thesis - unfolding frontier_def and closure_open_interval[OF False] and interior_open[OF open_interval] + unfolding frontier_def and closure_open_interval[OF False] and interior_open[OF open_box] by auto qed lemma inter_interval_mixed_eq_empty: - fixes a :: "'a::ordered_euclidean_space" - assumes "box c d \ {}" - shows "box a b \ {c .. d} = {} \ box a b \ box c d = {}" + fixes a :: "'a::euclidean_space" + assumes "box c d \ {}" + shows "box a b \ cbox c d = {} \ box a b \ box c d = {}" unfolding closure_open_interval[OF assms, symmetric] - unfolding open_inter_closure_eq_empty[OF open_interval] .. + unfolding open_inter_closure_eq_empty[OF open_box] .. lemma diameter_closed_interval: fixes a b::"'a::ordered_euclidean_space" @@ -726,9 +727,9 @@ definition "is_interval (s::('a::euclidean_space) set) \ (\a\s. \b\s. \x. (\i\Basis. ((a\i \ x\i \ x\i \ b\i) \ (b\i \ x\i \ x\i \ a\i))) \ x \ s)" -lemma is_interval_interval: "is_interval {a .. b::'a::ordered_euclidean_space}" (is ?th1) +lemma is_interval_interval: "is_interval (cbox a (b::'a::euclidean_space))" (is ?th1) "is_interval (box a b)" (is ?th2) proof - - show ?th1 ?th2 unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff + show ?th1 ?th2 unfolding is_interval_def mem_box Ball_def atLeastAtMost_iff by(meson order_trans le_less_trans less_le_trans less_trans)+ qed lemma is_interval_empty: @@ -756,7 +757,7 @@ shows "(\i\Basis. (if i = j then y i \ i else x \ i) *\<^sub>R i) \ S" by (rule mem_is_intervalI[OF assms(1,2)]) (auto simp: assms) -lemma mem_interval_componentwiseI: +lemma mem_box_componentwiseI: fixes S::"'a::ordered_euclidean_space set" assumes "is_interval S" assumes "\i. i \ Basis \ x \ i \ ((\x. x \ i) ` S)" @@ -787,93 +788,72 @@ finally show ?thesis . qed -text{* Instantiation for intervals on @{text ordered_euclidean_space} *} - -lemma eucl_lessThan_eq_halfspaces: - fixes a :: "'a\ordered_euclidean_space" +lemma eucl_less_eq_halfspaces: + fixes a :: "'a\euclidean_space" shows "{x. x i\Basis. {x. x \ i < a \ i})" - by (auto simp: eucl_less_def) - -lemma eucl_greaterThan_eq_halfspaces: - fixes a :: "'a\ordered_euclidean_space" - shows "{x. a i\Basis. {x. a \ i < x \ i})" + "{x. a i\Basis. {x. a \ i < x \ i})" by (auto simp: eucl_less_def) -lemma eucl_atMost_eq_halfspaces: - fixes a :: "'a\ordered_euclidean_space" - shows "{.. a} = (\i\Basis. {x. x \ i \ a \ i})" - by (auto simp: eucl_le[where 'a='a]) - -lemma eucl_atLeast_eq_halfspaces: - fixes a :: "'a\ordered_euclidean_space" - shows "{a ..} = (\i\Basis. {x. a \ i \ x \ i})" - by (auto simp: eucl_le[where 'a='a]) +lemma eucl_le_eq_halfspaces: + fixes a :: "'a\euclidean_space" + shows "{x. \i\Basis. x \ i \ a \ i} = (\i\Basis. {x. x \ i \ a \ i})" + "{x. \i\Basis. a \ i \ x \ i} = (\i\Basis. {x. a \ i \ x \ i})" + by auto -lemma open_eucl_lessThan[simp, intro]: - fixes a :: "'a\ordered_euclidean_space" +lemma open_Collect_eucl_less[simp, intro]: + fixes a :: "'a\euclidean_space" shows "open {x. x ordered_euclidean_space" - shows "open {x. a ordered_euclidean_space" - shows "closed {.. a}" - unfolding eucl_atMost_eq_halfspaces - by (simp add: closed_INT closed_Collect_le) +lemma closed_Collect_eucl_le[simp, intro]: + fixes a :: "'a\euclidean_space" + shows "closed {x. \i\Basis. a \ i \ x \ i}" + "closed {x. \i\Basis. x \ i \ a \ i}" + unfolding eucl_le_eq_halfspaces + by (simp_all add: closed_INT closed_Collect_le) -lemma closed_eucl_atLeast[simp, intro]: - fixes a :: "'a\ordered_euclidean_space" - shows "closed {a ..}" - unfolding eucl_atLeast_eq_halfspaces - by (simp add: closed_INT closed_Collect_le) - - -lemma image_affinity_interval: fixes m::real - fixes a b c :: "'a::ordered_euclidean_space" - shows "(\x. m *\<^sub>R x + c) ` {a .. b} = - (if {a .. b} = {} then {} - else (if 0 \ m then {m *\<^sub>R a + c .. m *\<^sub>R b + c} - else {m *\<^sub>R b + c .. m *\<^sub>R a + c}))" +lemma image_affinity_cbox: fixes m::real + fixes a b c :: "'a::euclidean_space" + shows "(\x. m *\<^sub>R x + c) ` cbox a b = + (if cbox a b = {} then {} + else (if 0 \ m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) + else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" proof (cases "m = 0") case True { fix x - assume "x \ c" "c \ x" + assume "\i\Basis. x \ i \ c \ i" "\i\Basis. c \ i \ x \ i" then have "x = c" - unfolding eucl_le[where 'a='a] apply - apply (subst euclidean_eq_iff) apply (auto intro: order_antisym) done } - moreover have "c \ {m *\<^sub>R a + c..m *\<^sub>R b + c}" - unfolding True by (auto simp add: eucl_le[where 'a='a]) - ultimately show ?thesis using True by auto + moreover have "c \ cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" + unfolding True by (auto simp add: cbox_sing) + ultimately show ?thesis using True by (auto simp: cbox_def) next case False { fix y - assume "a \ y" "y \ b" "m > 0" - then have "m *\<^sub>R a + c \ m *\<^sub>R y + c" and "m *\<^sub>R y + c \ m *\<^sub>R b + c" - unfolding eucl_le[where 'a='a] by (auto simp: inner_distrib) + assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m > 0" + then have "\i\Basis. (m *\<^sub>R a + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R b + c) \ i" + by (auto simp: inner_distrib) } moreover { fix y - assume "a \ y" "y \ b" "m < 0" - then have "m *\<^sub>R b + c \ m *\<^sub>R y + c" and "m *\<^sub>R y + c \ m *\<^sub>R a + c" - unfolding eucl_le[where 'a='a] by (auto simp add: mult_left_mono_neg inner_distrib) + assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m < 0" + then have "\i\Basis. (m *\<^sub>R b + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R a + c) \ i" + by (auto simp add: mult_left_mono_neg inner_distrib) } moreover { fix y - assume "m > 0" and "m *\<^sub>R a + c \ y" and "y \ m *\<^sub>R b + c" - then have "y \ (\x. m *\<^sub>R x + c) ` {a..b}" - unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a] + assume "m > 0" and "\i\Basis. (m *\<^sub>R a + c) \ i \ y \ i" and "\i\Basis. y \ i \ (m *\<^sub>R b + c) \ i" + then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" + unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left) done @@ -881,19 +861,58 @@ moreover { fix y - assume "m *\<^sub>R b + c \ y" "y \ m *\<^sub>R a + c" "m < 0" - then have "y \ (\x. m *\<^sub>R x + c) ` {a..b}" - unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a] + assume "\i\Basis. (m *\<^sub>R b + c) \ i \ y \ i" "\i\Basis. y \ i \ (m *\<^sub>R a + c) \ i" "m < 0" + then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" + unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left) done } - ultimately show ?thesis using False by auto + ultimately show ?thesis using False by (auto simp: cbox_def) qed -lemma image_smult_interval:"(\x. m *\<^sub>R (x::_::ordered_euclidean_space)) ` {a..b} = - (if {a..b} = {} then {} else if 0 \ m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R a})" - using image_affinity_interval[of m 0 a b] by auto +lemma image_smult_cbox:"(\x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = + (if cbox a b = {} then {} else if 0 \ m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))" + using image_affinity_cbox[of m 0 a b] by auto + +text{* Instantiation for intervals on @{text ordered_euclidean_space} *} + +lemma + fixes a :: "'a\ordered_euclidean_space" + shows cbox_interval: "cbox a b = {a..b}" + and interval_cbox: "{a..b} = cbox a b" + and eucl_le_atMost: "{x. \i\Basis. x \ i <= a \ i} = {..a}" + and eucl_le_atLeast: "{x. \i\Basis. a \ i <= x \ i} = {a..}" + by (auto simp: eucl_le[where 'a='a] eucl_less_def box_def cbox_def) + +lemma closed_eucl_atLeastAtMost[simp, intro]: + fixes a :: "'a\ordered_euclidean_space" + shows "closed {a..b}" + by (simp add: cbox_interval[symmetric] closed_cbox) + +lemma closed_eucl_atMost[simp, intro]: + fixes a :: "'a\ordered_euclidean_space" + shows "closed {..a}" + by (simp add: eucl_le_atMost[symmetric]) + +lemma closed_eucl_atLeast[simp, intro]: + fixes a :: "'a\ordered_euclidean_space" + shows "closed {a..}" + by (simp add: eucl_le_atLeast[symmetric]) + +lemma image_affinity_interval: fixes m::real + fixes a b c :: "'a::ordered_euclidean_space" + shows "(\x. m *\<^sub>R x + c) ` {a..b} = + (if {a..b} = {} then {} + else (if 0 \ m then {m *\<^sub>R a + c .. m *\<^sub>R b + c} + else {m *\<^sub>R b + c .. m *\<^sub>R a + c}))" + using image_affinity_cbox[of m c a b] + by (simp add: cbox_interval) + +lemma image_smult_interval:"(\x. m *\<^sub>R (x::_::ordered_euclidean_space)) ` {a .. b} = + (if {a .. b} = {} then {} else if 0 \ m then {m *\<^sub>R a .. m *\<^sub>R b} else {m *\<^sub>R b .. m *\<^sub>R a})" + using image_smult_cbox[of m a b] + by (simp add: cbox_interval) no_notation eucl_less (infix " {}" - unfolding path_image_def image_is_empty interval_eq_empty + unfolding path_image_def image_is_empty box_eq_empty by auto lemma pathstart_in_path_image[intro]: "pathstart g \ path_image g" @@ -64,7 +64,7 @@ lemma compact_path_image[intro]: "path g \ compact (path_image g)" unfolding path_def path_image_def apply (erule compact_continuous_image) - apply (rule compact_interval) + apply (rule compact_Icc) done lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g" diff -r 2666cd7d380c -r 0268784f60da src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Mar 17 21:56:32 2014 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Mar 18 10:12:57 2014 +0100 @@ -826,10 +826,23 @@ where "eucl_less a b \ (\i\Basis. a \ i < b \ i)" definition box_eucl_less: "box a b = {x. a x i\Basis. a \ i \ x \ i \ x \ i \ b \ i}" lemma box_def: "box a b = {x. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" and in_box_eucl_less: "x \ box a b \ a x box a b \ (\i\Basis. a \ i < x \ i \ x \ i < b \ i)" + "x \ cbox a b \ (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i)" + by (auto simp: box_eucl_less eucl_less_def cbox_def) + +lemma mem_box_real[simp]: + "(x::real) \ box a b \ a < x \ x < b" + "(x::real) \ cbox a b \ a \ x \ x \ b" + by (auto simp: mem_box) + +lemma box_real[simp]: + fixes a b:: real + shows "box a b = {a <..< b}" "cbox a b = {a .. b}" + by auto lemma rational_boxes: fixes x :: "'a\euclidean_space" diff -r 2666cd7d380c -r 0268784f60da src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Mon Mar 17 21:56:32 2014 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Tue Mar 18 10:12:57 2014 +0100 @@ -48,10 +48,10 @@ by (fastforce simp: eucl_le[where 'a='a] cube_def setsum_negf) lemma cube_subset_iff: "cube n \ cube N \ n \ N" - unfolding cube_def subset_interval by (simp add: setsum_negf ex_in_conv) + unfolding cube_def subset_box by (simp add: setsum_negf ex_in_conv eucl_le[where 'a='a]) lemma ball_subset_cube: "ball (0::'a::ordered_euclidean_space) (real n) \ cube n" - apply (simp add: cube_def subset_eq mem_interval setsum_negf eucl_le[where 'a='a]) + apply (simp add: cube_def subset_eq mem_box setsum_negf eucl_le[where 'a='a]) proof safe fix x i :: 'a assume x: "x \ ball 0 (real n)" and i: "i \ Basis" thus "- real n \ x \ i" "real n \ x \ i" @@ -67,7 +67,7 @@ qed lemma cube_subset_Suc[intro]: "cube n \ cube (Suc n)" - unfolding cube_def subset_interval by (simp add: setsum_negf) + unfolding cube_def cbox_interval[symmetric] subset_box by (simp add: setsum_negf) lemma has_integral_interval_cube: fixes a b :: "'a::ordered_euclidean_space" @@ -81,7 +81,7 @@ also have "\ \ ((\x. 1::real) has_integral content ?R *\<^sub>R 1) ?R" unfolding indicator_def [abs_def] has_integral_restrict_univ real_scaleR_def mult_1_right .. also have "((\x. 1) has_integral content ?R *\<^sub>R 1) ?R" - unfolding cube_def inter_interval by (rule has_integral_const) + unfolding cube_def inter_interval cbox_interval[symmetric] by (rule has_integral_const) finally show ?thesis . qed @@ -222,7 +222,7 @@ lemma lebesgueI_negligible[dest]: fixes s::"'a::ordered_euclidean_space set" assumes "negligible s" shows "s \ sets lebesgue" - using assms by (force simp: cube_def integrable_on_def negligible_def intro!: lebesgueI) + using assms by (force simp: cbox_interval[symmetric] cube_def integrable_on_def negligible_def intro!: lebesgueI) lemma lmeasure_eq_0: fixes S :: "'a::ordered_euclidean_space set" @@ -231,7 +231,7 @@ have "\n. integral (cube n) (indicator S :: 'a\real) = 0" unfolding lebesgue_integral_def using assms by (intro integral_unique some1_equality ex_ex1I) - (auto simp: cube_def negligible_def) + (auto simp: cube_def negligible_def cbox_interval[symmetric]) then show ?thesis using assms by (simp add: emeasure_lebesgue lebesgueI_negligible) qed @@ -381,15 +381,15 @@ show "negligible A" unfolding negligible_def proof (intro allI) fix a b :: 'a - have integrable: "(indicator A :: _\real) integrable_on {a..b}" - by (intro integrable_on_subinterval has_integral_integrable) (auto intro: *) - then have "integral {a..b} (indicator A) \ (integral UNIV (indicator A) :: real)" + have integrable: "(indicator A :: _\real) integrable_on cbox a b" + by (intro integrable_on_subcbox has_integral_integrable) (auto intro: *) + then have "integral (cbox a b) (indicator A) \ (integral UNIV (indicator A) :: real)" using * by (auto intro!: integral_subset_le) - moreover have "(0::real) \ integral {a..b} (indicator A)" + moreover have "(0::real) \ integral (cbox a b) (indicator A)" using integrable by (auto intro!: integral_nonneg) - ultimately have "integral {a..b} (indicator A) = (0::real)" + ultimately have "integral (cbox a b) (indicator A) = (0::real)" using integral_unique[OF *] by auto - then show "(indicator A has_integral (0::real)) {a..b}" + then show "(indicator A has_integral (0::real)) (cbox a b)" using integrable_integral[OF integrable] by simp qed qed @@ -412,7 +412,7 @@ qed } ultimately show "ereal (real n) \ ereal (integral (cube n) (indicator UNIV::'a\real))" using integral_const DIM_positive[where 'a='a] - by (auto simp: cube_def content_closed_interval_cases setprod_constant setsum_negf) + by (auto simp: cube_def content_cbox_cases setprod_constant setsum_negf cbox_interval[symmetric]) qed simp lemma lmeasure_complete: "A \ B \ B \ null_sets lebesgue \ A \ null_sets lebesgue" @@ -423,9 +423,9 @@ shows lmeasure_atLeastAtMost[simp]: "emeasure lebesgue {a..b} = ereal (content {a..b})" proof - have "(indicator (UNIV \ {a..b})::_\real) integrable_on UNIV" - unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def]) + unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def] cbox_interval[symmetric]) from lmeasure_eq_integral[OF this] show ?thesis unfolding integral_indicator_UNIV - by (simp add: indicator_def [abs_def]) + by (simp add: indicator_def [abs_def] cbox_interval[symmetric]) qed lemma lmeasure_singleton[simp]: @@ -505,7 +505,7 @@ lemma Int_stable_atLeastAtMost: fixes x::"'a::ordered_euclidean_space" shows "Int_stable (range (\(a, b::'a). {a..b}))" - by (auto simp: inter_interval Int_stable_def) + by (auto simp: inter_interval Int_stable_def cbox_interval[symmetric]) lemma lborel_eqI: fixes M :: "'a::ordered_euclidean_space measure" @@ -959,9 +959,9 @@ proof cases assume "{a..b} \ {}" then have "a \ b" - by (simp add: interval_ne_empty eucl_le[where 'a='a]) + by (simp add: eucl_le[where 'a='a]) then have "emeasure lborel {a..b} = (\x\Basis. emeasure lborel {a \ x .. b \ x})" - by (auto simp: content_closed_interval eucl_le[where 'a='a] + by (auto simp: eucl_le[where 'a='a] content_closed_interval intro!: setprod_ereal[symmetric]) also have "\ = emeasure ?P (p2e -` {a..b} \ space ?P)" unfolding * by (subst lborel_space.measure_times) auto