Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
--- a/NEWS Fri Dec 14 14:46:01 2012 +0100
+++ b/NEWS Fri Dec 14 15:46:01 2012 +0100
@@ -227,6 +227,36 @@
* HOL/Cardinals: Theories of ordinals and cardinals
(supersedes the AFP entry "Ordinals_and_Cardinals").
+* HOL/Multivariate_Analysis:
+ Replaced "basis :: 'a::euclidean_space => nat => real" and
+ "\<Chi>\<Chi> :: (nat => real) => 'a::euclidean_space" on euclidean spaces by
+ using the inner product "_ \<bullet> _" with vectors from the Basis set.
+ "\<Chi>\<Chi> i. f i" is replaced by "SUM i : Basis. f i *r i".
+
+ With this change the following constants are also chanegd or removed:
+
+ DIM('a) :: nat ~> card (Basis :: 'a set) (is an abbreviation)
+ a $$ i ~> inner a i (where i : Basis)
+ cart_base i removed
+ \<pi>, \<pi>' removed
+
+ Theorems about these constants where removed.
+
+ Renamed lemmas:
+
+ component_le_norm ~> Basis_le_norm
+ euclidean_eq ~> euclidean_eq_iff
+ differential_zero_maxmin_component ~> differential_zero_maxmin_cart
+ euclidean_simps ~> inner_simps
+ independent_basis ~> independent_Basis
+ span_basis ~> span_Basis
+ in_span_basis ~> in_span_Basis
+ norm_bound_component_le ~> norm_boound_Basis_le
+ norm_bound_component_lt ~> norm_boound_Basis_lt
+ component_le_infnorm ~> Basis_le_infnorm
+
+ INCOMPATIBILITY.
+
* HOL/Probability:
- Add simproc "measurable" to automatically prove measurability
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Fri Dec 14 15:46:01 2012 +0100
@@ -22,6 +22,18 @@
imports Convex_Euclidean_Space
begin
+(** move this **)
+lemma divide_nonneg_nonneg:assumes "a \<ge> 0" "b \<ge> 0" shows "0 \<le> a / (b::real)"
+ apply(cases "b=0") defer apply(rule divide_nonneg_pos) using assms by auto
+
+lemma continuous_setsum:
+ fixes f :: "'i \<Rightarrow> 'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+ assumes f: "\<And>i. i \<in> I \<Longrightarrow> continuous F (f i)" shows "continuous F (\<lambda>x. \<Sum>i\<in>I. f i x)"
+proof cases
+ assume "finite I" from this f show ?thesis
+ by (induct I) (auto intro!: continuous_intros)
+qed (auto intro!: continuous_intros)
+
lemma brouwer_compactness_lemma:
assumes "compact s" "continuous_on s f" "\<not> (\<exists>x\<in>s. (f x = (0::_::euclidean_space)))"
obtains d where "0 < d" "\<forall>x\<in>s. d \<le> norm(f x)"
@@ -39,39 +51,39 @@
qed
lemma kuhn_labelling_lemma:
- fixes type::"'a::euclidean_space"
- assumes "(\<forall>x::'a. P x \<longrightarrow> P (f x))"
- and "\<forall>x. P x \<longrightarrow> (\<forall>i<DIM('a). Q i \<longrightarrow> 0 \<le> x$$i \<and> x$$i \<le> 1)"
- shows "\<exists>l. (\<forall>x.\<forall> i<DIM('a). l x i \<le> (1::nat)) \<and>
- (\<forall>x.\<forall> i<DIM('a). P x \<and> Q i \<and> (x$$i = 0) \<longrightarrow> (l x i = 0)) \<and>
- (\<forall>x.\<forall> i<DIM('a). P x \<and> Q i \<and> (x$$i = 1) \<longrightarrow> (l x i = 1)) \<and>
- (\<forall>x.\<forall> i<DIM('a). P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x$$i \<le> f(x)$$i) \<and>
- (\<forall>x.\<forall> i<DIM('a). P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f(x)$$i \<le> x$$i)"
+ fixes P Q :: "'a::euclidean_space \<Rightarrow> bool"
+ assumes "(\<forall>x. P x \<longrightarrow> P (f x))"
+ and "\<forall>x. P x \<longrightarrow> (\<forall>i\<in>Basis. Q i \<longrightarrow> 0 \<le> x\<bullet>i \<and> x\<bullet>i \<le> 1)"
+ shows "\<exists>l. (\<forall>x.\<forall>i\<in>Basis. l x i \<le> (1::nat)) \<and>
+ (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 0) \<longrightarrow> (l x i = 0)) \<and>
+ (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 1) \<longrightarrow> (l x i = 1)) \<and>
+ (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x\<bullet>i \<le> f(x)\<bullet>i) \<and>
+ (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f(x)\<bullet>i \<le> x\<bullet>i)"
proof -
have and_forall_thm:"\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)"
by auto
have *: "\<forall>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 \<longrightarrow> (x \<noteq> 1 \<and> x \<le> y \<or> x \<noteq> 0 \<and> y \<le> x)"
by auto
show ?thesis
- unfolding and_forall_thm
+ unfolding and_forall_thm Ball_def
apply(subst choice_iff[THEN sym])+
apply rule
apply rule
proof -
- case goal1
- let ?R = "\<lambda>y. (P x \<and> Q xa \<and> x $$ xa = 0 \<longrightarrow> y = (0::nat)) \<and>
- (P x \<and> Q xa \<and> x $$ xa = 1 \<longrightarrow> y = 1) \<and>
- (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x $$ xa \<le> f x $$ xa) \<and>
- (P x \<and> Q xa \<and> y = 1 \<longrightarrow> f x $$ xa \<le> x $$ xa)"
+ case (goal1 x)
+ let ?R = "\<lambda>y. (P x \<and> Q xa \<and> x \<bullet> xa = 0 \<longrightarrow> y = (0::nat)) \<and>
+ (P x \<and> Q xa \<and> x \<bullet> xa = 1 \<longrightarrow> y = 1) \<and>
+ (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x \<bullet> xa \<le> f x \<bullet> xa) \<and>
+ (P x \<and> Q xa \<and> y = 1 \<longrightarrow> f x \<bullet> xa \<le> x \<bullet> xa)"
{
- assume "P x" "Q xa" "xa<DIM('a)"
- then have "0 \<le> f x $$ xa \<and> f x $$ xa \<le> 1"
+ assume "P x" "Q xa" "xa\<in>Basis"
+ then have "0 \<le> f x \<bullet> xa \<and> f x \<bullet> xa \<le> 1"
using assms(2)[rule_format,of "f x" xa]
apply (drule_tac assms(1)[rule_format])
apply auto
done
}
- then have "xa<DIM('a) \<Longrightarrow> ?R 0 \<or> ?R 1" by auto
+ then have "xa\<in>Basis \<Longrightarrow> ?R 0 \<or> ?R 1" by auto
then show ?case by auto
qed
qed
@@ -1363,50 +1375,56 @@
apply(drule_tac assms(1)[rule_format]) by auto }
hence "?R 0 \<or> ?R 1" by auto thus ?case by auto qed qed
-lemma brouwer_cube: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space"
- assumes "continuous_on {0..\<chi>\<chi> i. 1} f" "f ` {0..\<chi>\<chi> i. 1} \<subseteq> {0..\<chi>\<chi> i. 1}"
- shows "\<exists>x\<in>{0..\<chi>\<chi> i. 1}. f x = x" apply(rule ccontr) proof-
- def n \<equiv> "DIM('a)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_def by(auto simp add:Suc_le_eq)
- assume "\<not> (\<exists>x\<in>{0..\<chi>\<chi> i. 1}. f x = x)" hence *:"\<not> (\<exists>x\<in>{0..\<chi>\<chi> i. 1}. f x - x = 0)" by auto
+lemma brouwer_cube:
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space"
+ assumes "continuous_on {0..(\<Sum>Basis)} f" "f ` {0..(\<Sum>Basis)} \<subseteq> {0..(\<Sum>Basis)}"
+ shows "\<exists>x\<in>{0..(\<Sum>Basis)}. f x = x"
+ proof (rule ccontr)
+ def n \<equiv> "DIM('a)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_def by(auto simp add: Suc_le_eq DIM_positive)
+ assume "\<not> (\<exists>x\<in>{0..\<Sum>Basis}. f x = x)" hence *:"\<not> (\<exists>x\<in>{0..\<Sum>Basis}. f x - x = 0)" by auto
guess d apply(rule brouwer_compactness_lemma[OF compact_interval _ *])
apply(rule continuous_on_intros assms)+ . note d=this[rule_format]
- have *:"\<forall>x. x \<in> {0..\<chi>\<chi> i. 1} \<longrightarrow> f x \<in> {0..\<chi>\<chi> i. 1}" "\<forall>x. x \<in> {0..(\<chi>\<chi> i. 1)::'a} \<longrightarrow>
- (\<forall>i<DIM('a). True \<longrightarrow> 0 \<le> x $$ i \<and> x $$ i \<le> 1)"
+ have *:"\<forall>x. x \<in> {0..\<Sum>Basis} \<longrightarrow> f x \<in> {0..\<Sum>Basis}" "\<forall>x. x \<in> {0..(\<Sum>Basis)::'a} \<longrightarrow>
+ (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)"
using assms(2)[unfolded image_subset_iff Ball_def] unfolding mem_interval by auto
guess label using kuhn_labelling_lemma[OF *] apply-apply(erule exE,(erule conjE)+) . note label = this[rule_format]
- have lem1:"\<forall>x\<in>{0..\<chi>\<chi> i. 1}.\<forall>y\<in>{0..\<chi>\<chi> i. 1}.\<forall>i<DIM('a). label x i \<noteq> label y i
- \<longrightarrow> abs(f x $$ i - x $$ i) \<le> norm(f y - f x) + norm(y - x)" proof safe
- fix x y::'a assume xy:"x\<in>{0..\<chi>\<chi> i. 1}" "y\<in>{0..\<chi>\<chi> i. 1}" fix i assume i:"label x i \<noteq> label y i" "i<DIM('a)"
+ have lem1:"\<forall>x\<in>{0..\<Sum>Basis}.\<forall>y\<in>{0..\<Sum>Basis}.\<forall>i\<in>Basis. label x i \<noteq> label y i
+ \<longrightarrow> abs(f x \<bullet> i - x \<bullet> i) \<le> norm(f y - f x) + norm(y - x)" proof safe
+ fix x y::'a assume xy:"x\<in>{0..\<Sum>Basis}" "y\<in>{0..\<Sum>Basis}"
+ fix i assume i:"label x i \<noteq> label y i" "i\<in>Basis"
have *:"\<And>x y fx fy::real. (x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy)
\<Longrightarrow> abs(fx - x) \<le> abs(fy - fx) + abs(y - x)" by auto
- have "\<bar>(f x - x) $$ i\<bar> \<le> abs((f y - f x)$$i) + abs((y - x)$$i)" unfolding euclidean_simps
+ have "\<bar>(f x - x) \<bullet> i\<bar> \<le> abs((f y - f x)\<bullet>i) + abs((y - x)\<bullet>i)"
+ unfolding inner_simps
apply(rule *) apply(cases "label x i = 0") apply(rule disjI1,rule) prefer 3 proof(rule disjI2,rule)
assume lx:"label x i = 0" hence ly:"label y i = 1" using i label(1)[of i y] by auto
- show "x $$ i \<le> f x $$ i" apply(rule label(4)[rule_format]) using xy lx i(2) by auto
- show "f y $$ i \<le> y $$ i" apply(rule label(5)[rule_format]) using xy ly i(2) by auto next
+ show "x \<bullet> i \<le> f x \<bullet> i" apply(rule label(4)[rule_format]) using xy lx i(2) by auto
+ show "f y \<bullet> i \<le> y \<bullet> i" apply(rule label(5)[rule_format]) using xy ly i(2) by auto next
assume "label x i \<noteq> 0" hence l:"label x i = 1" "label y i = 0"
using i label(1)[of i x] label(1)[of i y] by auto
- show "f x $$ i \<le> x $$ i" apply(rule label(5)[rule_format]) using xy l i(2) by auto
- show "y $$ i \<le> f y $$ i" apply(rule label(4)[rule_format]) using xy l i(2) by auto qed
- also have "\<dots> \<le> norm (f y - f x) + norm (y - x)" apply(rule add_mono) by(rule component_le_norm)+
- finally show "\<bar>f x $$ i - x $$ i\<bar> \<le> norm (f y - f x) + norm (y - x)" unfolding euclidean_simps . qed
- have "\<exists>e>0. \<forall>x\<in>{0..\<chi>\<chi> i. 1}. \<forall>y\<in>{0..\<chi>\<chi> i. 1}. \<forall>z\<in>{0..\<chi>\<chi> i. 1}. \<forall>i<DIM('a).
- norm(x - z) < e \<and> norm(y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow> abs((f(z) - z)$$i) < d / (real n)" proof-
- have d':"d / real n / 8 > 0" apply(rule divide_pos_pos)+ using d(1) unfolding n_def by auto
- have *:"uniformly_continuous_on {0..\<chi>\<chi> i. 1} f" by(rule compact_uniformly_continuous[OF assms(1) compact_interval])
+ show "f x \<bullet> i \<le> x \<bullet> i" apply(rule label(5)[rule_format]) using xy l i(2) by auto
+ show "y \<bullet> i \<le> f y \<bullet> i" apply(rule label(4)[rule_format]) using xy l i(2) by auto qed
+ also have "\<dots> \<le> norm (f y - f x) + norm (y - x)" apply(rule add_mono) by(rule Basis_le_norm[OF i(2)])+
+ finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" unfolding inner_simps .
+ qed
+ have "\<exists>e>0. \<forall>x\<in>{0..\<Sum>Basis}. \<forall>y\<in>{0..\<Sum>Basis}. \<forall>z\<in>{0..\<Sum>Basis}. \<forall>i\<in>Basis.
+ norm(x - z) < e \<and> norm(y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow> abs((f(z) - z)\<bullet>i) < d / (real n)" proof-
+ have d':"d / real n / 8 > 0" apply(rule divide_pos_pos)+ using d(1) unfolding n_def by (auto simp: DIM_positive)
+ have *:"uniformly_continuous_on {0..\<Sum>Basis} f" by(rule compact_uniformly_continuous[OF assms(1) compact_interval])
guess e using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] apply-apply(erule exE,(erule conjE)+) .
note e=this[rule_format,unfolded dist_norm]
show ?thesis apply(rule_tac x="min (e/2) (d/real n/8)" in exI)
- proof safe show "0 < min (e / 2) (d / real n / 8)" using d' e by auto
- fix x y z i assume as:"x \<in> {0..\<chi>\<chi> i. 1}" "y \<in> {0..\<chi>\<chi> i. 1}" "z \<in> {0..\<chi>\<chi> i. 1}"
+ proof safe
+ show "0 < min (e / 2) (d / real n / 8)" using d' e by auto
+ fix x y z i assume as:"x \<in> {0..\<Sum>Basis}" "y \<in> {0..\<Sum>Basis}" "z \<in> {0..\<Sum>Basis}"
"norm (x - z) < min (e / 2) (d / real n / 8)"
- "norm (y - z) < min (e / 2) (d / real n / 8)" "label x i \<noteq> label y i" and i:"i<DIM('a)"
+ "norm (y - z) < min (e / 2) (d / real n / 8)" "label x i \<noteq> label y i" and i:"i\<in>Basis"
have *:"\<And>z fz x fx n1 n2 n3 n4 d4 d::real. abs(fx - x) \<le> n1 + n2 \<Longrightarrow> abs(fx - fz) \<le> n3 \<Longrightarrow> abs(x - z) \<le> n4 \<Longrightarrow>
n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow> (8 * d4 = d) \<Longrightarrow> abs(fz - z) < d" by auto
- show "\<bar>(f z - z) $$ i\<bar> < d / real n" unfolding euclidean_simps proof(rule *)
- show "\<bar>f x $$ i - x $$ i\<bar> \<le> norm (f y -f x) + norm (y - x)" apply(rule lem1[rule_format]) using as i by auto
- show "\<bar>f x $$ i - f z $$ i\<bar> \<le> norm (f x - f z)" "\<bar>x $$ i - z $$ i\<bar> \<le> norm (x - z)"
- unfolding euclidean_component_diff[THEN sym] by(rule component_le_norm)+
+ show "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" unfolding inner_simps proof(rule *)
+ show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y -f x) + norm (y - x)" apply(rule lem1[rule_format]) using as i by auto
+ show "\<bar>f x \<bullet> i - f z \<bullet> i\<bar> \<le> norm (f x - f z)" "\<bar>x \<bullet> i - z \<bullet> i\<bar> \<le> norm (x - z)"
+ unfolding inner_diff_left[THEN sym] by(rule Basis_le_norm[OF i])+
have tria:"norm (y - x) \<le> norm (y - z) + norm (x - z)" using dist_triangle[of y x z,unfolded dist_norm]
unfolding norm_minus_commute by auto
also have "\<dots> < e / 2 + e / 2" apply(rule add_strict_mono) using as(4,5) by auto
@@ -1418,95 +1436,99 @@
guess p using real_arch_simple[of "1 + real n / e"] .. note p=this
have "1 + real n / e > 0" apply(rule add_pos_pos) defer apply(rule divide_pos_pos) using e(1) n by auto
hence "p > 0" using p by auto
- def b \<equiv> "\<lambda>i. i - 1::nat" have b:"bij_betw b {1..n} {..<DIM('a)}"
- unfolding bij_betw_def inj_on_def b_def n_def apply rule defer
- apply safe defer unfolding image_iff apply(rule_tac x="Suc x" in bexI) by auto
+
+ obtain b :: "nat \<Rightarrow> 'a" where b: "bij_betw b {1..n} Basis"
+ by atomize_elim (auto simp: n_def intro!: finite_same_card_bij)
def b' \<equiv> "inv_into {1..n} b"
- have b':"bij_betw b' {..<DIM('a)} {1..n}" using bij_betw_inv_into[OF b] unfolding b'_def n_def by auto
- have bb'[simp]:"\<And>i. i<DIM('a) \<Longrightarrow> b (b' i) = i" unfolding b'_def apply(rule f_inv_into_f) using b
- unfolding bij_betw_def by auto
- have b'b[simp]:"\<And>i. i\<in>{1..n} \<Longrightarrow> b' (b i) = i" unfolding b'_def apply(rule inv_into_f_eq)
- using b unfolding n_def bij_betw_def by auto
+ then have b': "bij_betw b' Basis {1..n}"
+ using bij_betw_inv_into[OF b] by auto
+ then have b'_Basis: "\<And>i. i \<in> Basis \<Longrightarrow> b' i \<in> {Suc 0 .. n}"
+ unfolding bij_betw_def by (auto simp: set_eq_iff)
+ have bb'[simp]:"\<And>i. i \<in> Basis \<Longrightarrow> b (b' i) = i"
+ unfolding b'_def using b by (auto simp: f_inv_into_f bij_betw_def)
+ have b'b[simp]:"\<And>i. i \<in> {1..n} \<Longrightarrow> b' (b i) = i"
+ unfolding b'_def using b by (auto simp: inv_into_f_eq bij_betw_def)
have *:"\<And>x::nat. x=0 \<or> x=1 \<longleftrightarrow> x\<le>1" by auto
- have b'':"\<And>j. j\<in>{1..n} \<Longrightarrow> b j <DIM('a)" using b unfolding bij_betw_def by auto
+ have b'':"\<And>j. j\<in>{Suc 0..n} \<Longrightarrow> b j \<in>Basis" using b unfolding bij_betw_def by auto
have q1:"0 < p" "0 < n" "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow>
- (\<forall>i\<in>{1..n}. (label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0 \<or> (label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1)"
+ (\<forall>i\<in>{1..n}. (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0 \<or>
+ (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)"
unfolding * using `p>0` `n>0` using label(1)[OF b''] by auto
- have q2:"\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = 0 \<longrightarrow> (label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0)"
- "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = p \<longrightarrow> (label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1)"
- apply(rule,rule,rule,rule) defer proof(rule,rule,rule,rule) fix x i
- assume as:"\<forall>i\<in>{1..n}. x i \<le> p" "i \<in> {1..n}"
- { assume "x i = p \<or> x i = 0"
- have "(\<chi>\<chi> i. real (x (b' i)) / real p) \<in> {0::'a..\<chi>\<chi> i. 1}" unfolding mem_interval
- apply safe unfolding euclidean_lambda_beta euclidean_component_zero
- proof (simp_all only: if_P) fix j assume j':"j<DIM('a)"
- hence j:"b' j \<in> {1..n}" using b' unfolding n_def bij_betw_def by auto
- show "0 \<le> real (x (b' j)) / real p"
- apply(rule divide_nonneg_pos) using `p>0` using as(1)[rule_format,OF j] by auto
- show "real (x (b' j)) / real p \<le> 1" unfolding divide_le_eq_1
- using as(1)[rule_format,OF j] by auto qed } note cube=this
- { assume "x i = p" thus "(label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1" unfolding o_def
- apply- apply(rule label(3)) apply(rule b'') using cube using as `p>0`
- proof safe assume i:"i\<in>{1..n}"
- show "((\<chi>\<chi> ia. real (x (b' ia)) / real (x i))::'a) $$ b i = 1"
- unfolding euclidean_lambda_beta apply(subst if_P) apply(rule b''[OF i]) unfolding b'b[OF i]
- unfolding `x i = p` using q1(1) by auto
- qed auto }
- { assume "x i = 0" thus "(label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0" unfolding o_def
- apply-apply(rule label(2)[OF b'']) using cube using as `p>0`
- proof safe assume i:"i\<in>{1..n}"
- show "((\<chi>\<chi> ia. real (x (b' ia)) / real p)::'a) $$ b i = 0"
- unfolding euclidean_lambda_beta apply (subst if_P) apply(rule b''[OF i]) unfolding b'b[OF i]
- unfolding `x i = 0` using q1(1) by auto
- qed auto }
+ have q2:"\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = 0 \<longrightarrow>
+ (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)"
+ "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = p \<longrightarrow>
+ (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)"
+ apply(rule,rule,rule,rule)
+ defer
+ proof(rule,rule,rule,rule)
+ fix x i assume as:"\<forall>i\<in>{1..n}. x i \<le> p" "i \<in> {1..n}"
+ { assume "x i = p \<or> x i = 0"
+ have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> {0::'a..\<Sum>Basis}"
+ unfolding mem_interval using as b'_Basis
+ by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) }
+ note cube=this
+ { assume "x i = p" thus "(label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1"
+ unfolding o_def using cube as `p>0`
+ by (intro label(3)) (auto simp add: b'') }
+ { assume "x i = 0" thus "(label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0"
+ unfolding o_def using cube as `p>0`
+ by (intro label(2)) (auto simp add: b'') }
qed
guess q apply(rule kuhn_lemma[OF q1 q2]) . note q=this
- def z \<equiv> "(\<chi>\<chi> i. real (q (b' i)) / real p)::'a"
- have "\<exists>i<DIM('a). d / real n \<le> abs((f z - z)$$i)" proof(rule ccontr)
- have "\<forall>i<DIM('a). q (b' i) \<in> {0..<p}" using q(1) b'[unfolded bij_betw_def] by auto
- hence "\<forall>i<DIM('a). q (b' i) \<in> {0..p}" apply-apply(rule,erule_tac x=i in allE) by auto
- hence "z\<in>{0..\<chi>\<chi> i.1}" unfolding z_def mem_interval apply safe unfolding euclidean_lambda_beta
- unfolding euclidean_component_zero apply (simp_all only: if_P)
- apply(rule divide_nonneg_pos) using `p>0` unfolding divide_le_eq_1 by auto
- hence d_fz_z:"d \<le> norm (f z - z)" apply(drule_tac d) .
- case goal1 hence as:"\<forall>i<DIM('a). \<bar>f z $$ i - z $$ i\<bar> < d / real n" using `n>0` by(auto simp add:not_le)
- have "norm (f z - z) \<le> (\<Sum>i<DIM('a). \<bar>f z $$ i - z $$ i\<bar>)" unfolding euclidean_component_diff[THEN sym] by(rule norm_le_l1)
- also have "\<dots> < (\<Sum>i<DIM('a). d / real n)" apply(rule setsum_strict_mono) using as by auto
- also have "\<dots> = d" unfolding real_eq_of_nat n_def using n using DIM_positive[where 'a='a] by auto
+ def z \<equiv> "(\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)::'a"
+ have "\<exists>i\<in>Basis. d / real n \<le> abs((f z - z)\<bullet>i)"
+ proof(rule ccontr)
+ have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}"
+ using q(1) b' by (auto intro: less_imp_le simp: bij_betw_def)
+ hence "z\<in>{0..\<Sum>Basis}"
+ unfolding z_def mem_interval using b'_Basis
+ by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1)
+ hence d_fz_z:"d \<le> norm (f z - z)" by (rule d)
+ case goal1
+ hence as:"\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n"
+ using `n>0` by(auto simp add: not_le inner_simps)
+ have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)"
+ unfolding inner_diff_left[symmetric] by(rule norm_le_l1)
+ also have "\<dots> < (\<Sum>(i::'a)\<in>Basis. d / real n)" apply(rule setsum_strict_mono) using as by auto
+ also have "\<dots> = d" using DIM_positive[where 'a='a] by (auto simp: real_eq_of_nat n_def)
finally show False using d_fz_z by auto qed then guess i .. note i=this
have *:"b' i \<in> {1..n}" using i using b'[unfolded bij_betw_def] by auto
guess r using q(2)[rule_format,OF *] .. then guess s apply-apply(erule exE,(erule conjE)+) . note rs=this[rule_format]
- have b'_im:"\<And>i. i<DIM('a) \<Longrightarrow> b' i \<in> {1..n}" using b' unfolding bij_betw_def by auto
- def r' \<equiv> "(\<chi>\<chi> i. real (r (b' i)) / real p)::'a"
- have "\<And>i. i<DIM('a) \<Longrightarrow> r (b' i) \<le> p" apply(rule order_trans) apply(rule rs(1)[OF b'_im,THEN conjunct2])
+ have b'_im:"\<And>i. i\<in>Basis \<Longrightarrow> b' i \<in> {1..n}" using b' unfolding bij_betw_def by auto
+ def r' \<equiv> "(\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)::'a"
+ have "\<And>i. i\<in>Basis \<Longrightarrow> r (b' i) \<le> p" apply(rule order_trans) apply(rule rs(1)[OF b'_im,THEN conjunct2])
using q(1)[rule_format,OF b'_im] by(auto simp add: Suc_le_eq)
- hence "r' \<in> {0..\<chi>\<chi> i. 1}" unfolding r'_def mem_interval apply safe unfolding euclidean_lambda_beta euclidean_component_zero
- apply (simp only: if_P)
- apply(rule divide_nonneg_pos) using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>0` by auto
- def s' \<equiv> "(\<chi>\<chi> i. real (s (b' i)) / real p)::'a"
- have "\<And>i. i<DIM('a) \<Longrightarrow> s (b' i) \<le> p" apply(rule order_trans) apply(rule rs(2)[OF b'_im,THEN conjunct2])
+ hence "r' \<in> {0..\<Sum>Basis}"
+ unfolding r'_def mem_interval using b'_Basis
+ by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1)
+ def s' \<equiv> "(\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)::'a"
+ have "\<And>i. i\<in>Basis \<Longrightarrow> s (b' i) \<le> p" apply(rule order_trans) apply(rule rs(2)[OF b'_im,THEN conjunct2])
using q(1)[rule_format,OF b'_im] by(auto simp add: Suc_le_eq)
- hence "s' \<in> {0..\<chi>\<chi> i.1}" unfolding s'_def mem_interval apply safe unfolding euclidean_lambda_beta euclidean_component_zero
- apply (simp_all only: if_P) apply(rule divide_nonneg_pos) using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>0` by auto
- have "z\<in>{0..\<chi>\<chi> i.1}" unfolding z_def mem_interval apply safe unfolding euclidean_lambda_beta euclidean_component_zero
- apply (simp_all only: if_P) apply(rule divide_nonneg_pos) using q(1)[rule_format,OF b'_im] `p>0` by(auto intro:less_imp_le)
+ hence "s' \<in> {0..\<Sum>Basis}"
+ unfolding s'_def mem_interval using b'_Basis
+ by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1)
+ have "z\<in>{0..\<Sum>Basis}"
+ unfolding z_def mem_interval using b'_Basis q(1)[rule_format,OF b'_im] `p>0`
+ by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le)
have *:"\<And>x. 1 + real x = real (Suc x)" by auto
- { have "(\<Sum>i<DIM('a). \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>i<DIM('a). 1)"
+ { have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
apply(rule setsum_mono) using rs(1)[OF b'_im] by(auto simp add:* field_simps)
- also have "\<dots> < e * real p" using p `e>0` `p>0` unfolding n_def real_of_nat_def
- by(auto simp add:field_simps)
- finally have "(\<Sum>i<DIM('a). \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . } moreover
- { have "(\<Sum>i<DIM('a). \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>i<DIM('a). 1)"
+ also have "\<dots> < e * real p" using p `e>0` `p>0`
+ by(auto simp add: field_simps n_def real_of_nat_def)
+ finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . } moreover
+ { have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
apply(rule setsum_mono) using rs(2)[OF b'_im] by(auto simp add:* field_simps)
- also have "\<dots> < e * real p" using p `e>0` `p>0` unfolding n_def real_of_nat_def
- by(auto simp add:field_simps)
- finally have "(\<Sum>i<DIM('a). \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . } ultimately
- have "norm (r' - z) < e" "norm (s' - z) < e" unfolding r'_def s'_def z_def apply-
- apply(rule_tac[!] le_less_trans[OF norm_le_l1]) using `p>0`
- by(auto simp add:field_simps setsum_divide_distrib[THEN sym])
- hence "\<bar>(f z - z) $$ i\<bar> < d / real n" apply-apply(rule e(2)[OF `r'\<in>{0..\<chi>\<chi> i.1}` `s'\<in>{0..\<chi>\<chi> i.1}` `z\<in>{0..\<chi>\<chi> i.1}`])
- using rs(3) unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' using i by auto
- thus False using i by auto qed
+ also have "\<dots> < e * real p" using p `e>0` `p>0`
+ by(auto simp add: field_simps n_def real_of_nat_def)
+ finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . } ultimately
+ have "norm (r' - z) < e" "norm (s' - z) < e" unfolding r'_def s'_def z_def using `p>0`
+ by (rule_tac[!] le_less_trans[OF norm_le_l1])
+ (auto simp add: field_simps setsum_divide_distrib[symmetric] inner_diff_left)
+ hence "\<bar>(f z - z) \<bullet> i\<bar> < d / real n"
+ using rs(3) i unfolding r'_def[symmetric] s'_def[symmetric] o_def bb'
+ by (intro e(2)[OF `r'\<in>{0..\<Sum>Basis}` `s'\<in>{0..\<Sum>Basis}` `z\<in>{0..\<Sum>Basis}`]) auto
+ thus False using i by auto
+qed
subsection {* Retractions. *}
@@ -1551,12 +1573,14 @@
subsection {*So the Brouwer theorem for any set with nonempty interior. *}
-lemma brouwer_weak: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space"
+lemma brouwer_weak:
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space"
assumes "compact s" "convex s" "interior s \<noteq> {}" "continuous_on s f" "f ` s \<subseteq> s"
obtains x where "x \<in> s" "f x = x" proof-
- have *:"interior {0::'a..\<chi>\<chi> i.1} \<noteq> {}" unfolding interior_closed_interval interval_eq_empty by auto
- have *:"{0::'a..\<chi>\<chi> i.1} homeomorphic s" using homeomorphic_convex_compact[OF convex_interval(1) compact_interval * assms(2,1,3)] .
- have "\<forall>f. continuous_on {0::'a..\<chi>\<chi> i.1} f \<and> f ` {0::'a..\<chi>\<chi> i.1} \<subseteq> {0::'a..\<chi>\<chi> i.1} \<longrightarrow> (\<exists>x\<in>{0::'a..\<chi>\<chi> i.1}. f x = x)"
+ have *:"interior {0::'a..\<Sum>Basis} \<noteq> {}" unfolding interior_closed_interval interval_eq_empty by auto
+ have *:"{0::'a..\<Sum>Basis} homeomorphic s" using homeomorphic_convex_compact[OF convex_interval(1) compact_interval * assms(2,1,3)] .
+ have "\<forall>f. continuous_on {0::'a..\<Sum>Basis} f \<and> f ` {0::'a..\<Sum>Basis} \<subseteq> {0::'a..\<Sum>Basis} \<longrightarrow>
+ (\<exists>x\<in>{0::'a..\<Sum>Basis}. f x = x)"
using brouwer_cube by auto
thus ?thesis unfolding homeomorphic_fixpoint_property[OF *] apply(erule_tac x=f in allE)
apply(erule impE) defer apply(erule bexE) apply(rule_tac x=y in that) using assms by auto qed
@@ -1609,49 +1633,45 @@
subsection {*Bijections between intervals. *}
-definition "interval_bij = (\<lambda> (a::'a,b::'a) (u::'a,v::'a) (x::'a::ordered_euclidean_space).
- (\<chi>\<chi> i. u$$i + (x$$i - a$$i) / (b$$i - a$$i) * (v$$i - u$$i))::'a)"
+definition interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::ordered_euclidean_space" where
+ "interval_bij \<equiv> \<lambda>(a, b) (u, v) x. (\<Sum>i\<in>Basis. (u\<bullet>i + (x\<bullet>i - a\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (v\<bullet>i - u\<bullet>i)) *\<^sub>R i)"
lemma interval_bij_affine:
- "interval_bij (a,b) (u,v) = (\<lambda>x. (\<chi>\<chi> i. (v$$i - u$$i) / (b$$i - a$$i) * x$$i) +
- (\<chi>\<chi> i. u$$i - (v$$i - u$$i) / (b$$i - a$$i) * a$$i))"
- apply rule apply(subst euclidean_eq,safe) unfolding euclidean_simps interval_bij_def euclidean_lambda_beta
- by(auto simp add: field_simps add_divide_distrib[THEN sym])
+ "interval_bij (a,b) (u,v) = (\<lambda>x. (\<Sum>i\<in>Basis. ((v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (x\<bullet>i)) *\<^sub>R i) +
+ (\<Sum>i\<in>Basis. (u\<bullet>i - (v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (a\<bullet>i)) *\<^sub>R i))"
+ by (auto simp: setsum_addf[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff
+ field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum_cong)
lemma continuous_interval_bij:
"continuous (at x) (interval_bij (a,b::'a::ordered_euclidean_space) (u,v::'a))"
- unfolding interval_bij_affine apply(rule continuous_intros)
- apply(rule linear_continuous_at) unfolding linear_conv_bounded_linear[THEN sym]
- unfolding linear_def euclidean_eq[where 'a='a] apply safe unfolding euclidean_lambda_beta prefer 3
- apply(rule continuous_intros) by(auto simp add:field_simps add_divide_distrib[THEN sym])
+ by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros)
lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a,b) (u,v))"
apply(rule continuous_at_imp_continuous_on) by(rule, rule continuous_interval_bij)
-(** move this **)
-lemma divide_nonneg_nonneg:assumes "a \<ge> 0" "b \<ge> 0" shows "0 \<le> a / (b::real)"
- apply(cases "b=0") defer apply(rule divide_nonneg_pos) using assms by auto
-
-lemma in_interval_interval_bij: assumes "x \<in> {a..b}" "{u..v} \<noteq> {}"
- shows "interval_bij (a,b) (u,v) x \<in> {u..v::'a::ordered_euclidean_space}"
- unfolding interval_bij_def split_conv mem_interval apply safe unfolding euclidean_lambda_beta
-proof (simp_all only: if_P)
- fix i assume i:"i<DIM('a)" have "{a..b} \<noteq> {}" using assms by auto
- hence *:"a$$i \<le> b$$i" "u$$i \<le> v$$i" using assms(2) unfolding interval_eq_empty not_ex apply-
- apply(erule_tac[!] x=i in allE)+ by auto
- have x:"a$$i\<le>x$$i" "x$$i\<le>b$$i" using assms(1)[unfolded mem_interval] using i by auto
- have "0 \<le> (x $$ i - a $$ i) / (b $$ i - a $$ i) * (v $$ i - u $$ i)"
- apply(rule mult_nonneg_nonneg) apply(rule divide_nonneg_nonneg)
- using * x by(auto simp add:field_simps)
- thus "u $$ i \<le> u $$ i + (x $$ i - a $$ i) / (b $$ i - a $$ i) * (v $$ i - u $$ i)" using * by auto
- have "((x $$ i - a $$ i) / (b $$ i - a $$ i)) * (v $$ i - u $$ i) \<le> 1 * (v $$ i - u $$ i)"
+lemma in_interval_interval_bij:
+ fixes a b u v x :: "'a::ordered_euclidean_space"
+ assumes "x \<in> {a..b}" "{u..v} \<noteq> {}" shows "interval_bij (a,b) (u,v) x \<in> {u..v}"
+ apply (simp only: interval_bij_def split_conv mem_interval inner_setsum_left_Basis cong: ball_cong)
+proof safe
+ fix i :: 'a assume i:"i\<in>Basis"
+ have "{a..b} \<noteq> {}" using assms by auto
+ with i have *: "a\<bullet>i \<le> b\<bullet>i" "u\<bullet>i \<le> v\<bullet>i"
+ using assms(2) by (auto simp add: interval_eq_empty not_less)
+ have x: "a\<bullet>i\<le>x\<bullet>i" "x\<bullet>i\<le>b\<bullet>i"
+ using assms(1)[unfolded mem_interval] using i by auto
+ have "0 \<le> (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)"
+ using * x by (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg)
+ thus "u \<bullet> i \<le> u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)"
+ using * by auto
+ have "((x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (v \<bullet> i - u \<bullet> i) \<le> 1 * (v \<bullet> i - u \<bullet> i)"
apply(rule mult_right_mono) unfolding divide_le_eq_1 using * x by auto
- thus "u $$ i + (x $$ i - a $$ i) / (b $$ i - a $$ i) * (v $$ i - u $$ i) \<le> v $$ i" using * by auto
+ thus "u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i) \<le> v \<bullet> i" using * by auto
qed
-lemma interval_bij_bij: fixes x::"'a::ordered_euclidean_space" assumes "\<forall>i. a$$i < b$$i \<and> u$$i < v$$i"
- shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x"
- unfolding interval_bij_def split_conv euclidean_eq[where 'a='a]
- apply(rule,insert assms,erule_tac x=i in allE) by auto
+lemma interval_bij_bij:
+ "\<forall>(i::'a::ordered_euclidean_space)\<in>Basis. a\<bullet>i < b\<bullet>i \<and> u\<bullet>i < v\<bullet>i \<Longrightarrow>
+ interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x"
+ by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a])
end
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Fri Dec 14 15:46:01 2012 +0100
@@ -328,283 +328,18 @@
fixes f:: "'a \<Rightarrow> real ^'n"
assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"
shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) * e"
-proof -
- let ?d = "real CARD('n)"
- let ?nf = "\<lambda>x. norm (f x)"
- let ?U = "UNIV :: 'n set"
- have th0: "setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $ i\<bar>) ?U) P = setsum (\<lambda>i. setsum (\<lambda>x. \<bar>f x $ i\<bar>) P) ?U"
- by (rule setsum_commute)
- have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def)
- have "setsum ?nf P \<le> setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $ i\<bar>) ?U) P"
- apply (rule setsum_mono)
- apply (rule norm_le_l1_cart)
- done
- also have "\<dots> \<le> 2 * ?d * e"
- unfolding th0 th1
- proof(rule setsum_bounded)
- fix i assume i: "i \<in> ?U"
- let ?Pp = "{x. x\<in> P \<and> f x $ i \<ge> 0}"
- let ?Pn = "{x. x \<in> P \<and> f x $ i < 0}"
- have thp: "P = ?Pp \<union> ?Pn" by auto
- have thp0: "?Pp \<inter> ?Pn ={}" by auto
- have PpP: "?Pp \<subseteq> P" and PnP: "?Pn \<subseteq> P" by blast+
- have Ppe:"setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp \<le> e"
- using component_le_norm_cart[of "setsum (\<lambda>x. f x) ?Pp" i] fPs[OF PpP]
- by (auto intro: abs_le_D1)
- have Pne: "setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn \<le> e"
- using component_le_norm_cart[of "setsum (\<lambda>x. - f x) ?Pn" i] fPs[OF PnP]
- by (auto simp add: setsum_negf intro: abs_le_D1)
- have "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P = setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp + setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn"
- apply (subst thp)
- apply (rule setsum_Un_zero)
- using fP thp0 apply auto
- done
- also have "\<dots> \<le> 2*e" using Pne Ppe by arith
- finally show "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P \<le> 2*e" .
- qed
- finally show ?thesis .
-qed
-
-lemma if_distr: "(if P then f else g) $ i = (if P then f $ i else g $ i)" by simp
-
-lemma split_dimensions'[consumes 1]:
- assumes "k < DIM('a::euclidean_space^'b)"
- obtains i j where "i < CARD('b::finite)"
- and "j < DIM('a::euclidean_space)"
- and "k = j + i * DIM('a::euclidean_space)"
- using split_times_into_modulo[OF assms[simplified]] .
-
-lemma cart_euclidean_bound[intro]:
- assumes j:"j < DIM('a::euclidean_space)"
- shows "j + \<pi>' (i::'b::finite) * DIM('a) < CARD('b) * DIM('a::euclidean_space)"
- using linear_less_than_times[OF pi'_range j, of i] .
-
-lemma (in euclidean_space) forall_CARD_DIM:
- "(\<forall>i<CARD('b) * DIM('a). P i) \<longleftrightarrow> (\<forall>(i::'b::finite) j. j<DIM('a) \<longrightarrow> P (j + \<pi>' i * DIM('a)))"
- (is "?l \<longleftrightarrow> ?r")
-proof (safe elim!: split_times_into_modulo)
- fix i :: 'b and j
- assume "j < DIM('a)"
- note linear_less_than_times[OF pi'_range[of i] this]
- moreover assume "?l"
- ultimately show "P (j + \<pi>' i * DIM('a))" by auto
-next
- fix i j
- assume "i < CARD('b)" "j < DIM('a)" and "?r"
- from `?r`[rule_format, OF `j < DIM('a)`, of "\<pi> i"] `i < CARD('b)`
- show "P (j + i * DIM('a))" by simp
-qed
-
-lemma (in euclidean_space) exists_CARD_DIM:
- "(\<exists>i<CARD('b) * DIM('a). P i) \<longleftrightarrow> (\<exists>i::'b::finite. \<exists>j<DIM('a). P (j + \<pi>' i * DIM('a)))"
- using forall_CARD_DIM[where 'b='b, of "\<lambda>x. \<not> P x"] by blast
-
-lemma forall_CARD:
- "(\<forall>i<CARD('b). P i) \<longleftrightarrow> (\<forall>i::'b::finite. P (\<pi>' i))"
- using forall_CARD_DIM[where 'a=real, of P] by simp
-
-lemma exists_CARD:
- "(\<exists>i<CARD('b). P i) \<longleftrightarrow> (\<exists>i::'b::finite. P (\<pi>' i))"
- using exists_CARD_DIM[where 'a=real, of P] by simp
-
-lemmas cart_simps = forall_CARD_DIM exists_CARD_DIM forall_CARD exists_CARD
-
-lemma cart_euclidean_nth[simp]:
- fixes x :: "('a::euclidean_space, 'b::finite) vec"
- assumes j:"j < DIM('a)"
- shows "x $$ (j + \<pi>' i * DIM('a)) = x $ i $$ j"
- unfolding euclidean_component_def inner_vec_def basis_eq_pi'[OF j] if_distrib cond_application_beta
- by (simp add: setsum_cases)
-
-lemma real_euclidean_nth:
- fixes x :: "real^'n"
- shows "x $$ \<pi>' i = (x $ i :: real)"
- using cart_euclidean_nth[where 'a=real, of 0 x i] by simp
-
-lemmas nth_conv_component = real_euclidean_nth[symmetric]
-
-lemma mult_split_eq:
- fixes A :: nat assumes "x < A" "y < A"
- shows "x + i * A = y + j * A \<longleftrightarrow> x = y \<and> i = j"
-proof
- assume *: "x + i * A = y + j * A"
- { fix x y i j assume "i < j" "x < A" and *: "x + i * A = y + j * A"
- hence "x + i * A < Suc i * A" using `x < A` by simp
- also have "\<dots> \<le> j * A" using `i < j` unfolding mult_le_cancel2 by simp
- also have "\<dots> \<le> y + j * A" by simp
- finally have "i = j" using * by simp }
- note eq = this
-
- have "i = j"
- proof (cases rule: linorder_cases)
- assume "i < j"
- from eq[OF this `x < A` *] show "i = j" by simp
- next
- assume "j < i"
- from eq[OF this `y < A` *[symmetric]] show "i = j" by simp
- qed simp
- thus "x = y \<and> i = j" using * by simp
-qed simp
+ using setsum_norm_allsubsets_bound[OF assms]
+ by (simp add: DIM_cart Basis_real_def)
instance vec :: (ordered_euclidean_space, finite) ordered_euclidean_space
proof
fix x y::"'a^'b"
- show "(x \<le> y) = (\<forall>i<DIM(('a, 'b) vec). x $$ i \<le> y $$ i)"
- unfolding less_eq_vec_def apply(subst eucl_le) by (simp add: cart_simps)
- show"(x < y) = (\<forall>i<DIM(('a, 'b) vec). x $$ i < y $$ i)"
- unfolding less_vec_def apply(subst eucl_less) by (simp add: cart_simps)
-qed
-
-
-subsection{* Basis vectors in coordinate directions. *}
-
-definition "cart_basis k = (\<chi> i. if i = k then 1 else 0)"
-
-lemma basis_component [simp]: "cart_basis k $ i = (if k=i then 1 else 0)"
- unfolding cart_basis_def by simp
-
-lemma norm_basis[simp]:
- shows "norm (cart_basis k :: real ^'n) = 1"
- apply (simp add: cart_basis_def norm_eq_sqrt_inner) unfolding inner_vec_def
- apply (vector delta_mult_idempotent)
- using setsum_delta[of "UNIV :: 'n set" "k" "\<lambda>k. 1::real"] apply auto
- done
-
-lemma norm_basis_1: "norm(cart_basis 1 :: real ^'n::{finite,one}) = 1"
- by (rule norm_basis)
-
-lemma vector_choose_size: "0 <= c ==> \<exists>(x::real^'n). norm x = c"
- by (rule exI[where x="c *\<^sub>R cart_basis arbitrary"]) simp
-
-lemma vector_choose_dist:
- assumes e: "0 <= e"
- shows "\<exists>(y::real^'n). dist x y = e"
-proof -
- from vector_choose_size[OF e] obtain c:: "real ^'n" where "norm c = e"
- by blast
- then have "dist x (x - c) = e" by (simp add: dist_norm)
- then show ?thesis by blast
+ show "(x \<le> y) = (\<forall>i\<in>Basis. x \<bullet> i \<le> y \<bullet> i)"
+ unfolding less_eq_vec_def apply(subst eucl_le) by (simp add: Basis_vec_def inner_axis)
+ show"(x < y) = (\<forall>i\<in>Basis. x \<bullet> i < y \<bullet> i)"
+ unfolding less_vec_def apply(subst eucl_less) by (simp add: Basis_vec_def inner_axis)
qed
-lemma basis_inj[intro]: "inj (cart_basis :: 'n \<Rightarrow> real ^'n)"
- by (simp add: inj_on_def vec_eq_iff)
-
-lemma basis_expansion: "setsum (\<lambda>i. (x$i) *s cart_basis i) UNIV = (x::('a::ring_1) ^'n)"
- (is "?lhs = ?rhs" is "setsum ?f ?S = _")
- by (auto simp add: vec_eq_iff
- if_distrib setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong)
-
-lemma smult_conv_scaleR: "c *s x = scaleR c x"
- unfolding vector_scalar_mult_def scaleR_vec_def by simp
-
-lemma basis_expansion': "setsum (\<lambda>i. (x$i) *\<^sub>R cart_basis i) UNIV = x"
- by (rule basis_expansion [where 'a=real, unfolded smult_conv_scaleR])
-
-lemma basis_expansion_unique:
- "setsum (\<lambda>i. f i *s cart_basis i) UNIV = (x::('a::comm_ring_1) ^'n) \<longleftrightarrow> (\<forall>i. f i = x$i)"
- by (simp add: vec_eq_iff setsum_delta if_distrib cong del: if_weak_cong)
-
-lemma dot_basis: "cart_basis i \<bullet> x = x$i" "x \<bullet> (cart_basis i) = (x$i)"
- by (auto simp add: inner_vec_def cart_basis_def cond_application_beta if_distrib setsum_delta
- cong del: if_weak_cong)
-
-lemma inner_basis:
- fixes x :: "'a::{real_inner, real_algebra_1} ^ 'n"
- shows "inner (cart_basis i) x = inner 1 (x $ i)"
- and "inner x (cart_basis i) = inner (x $ i) 1"
- unfolding inner_vec_def cart_basis_def
- by (auto simp add: cond_application_beta if_distrib setsum_delta cong del: if_weak_cong)
-
-lemma basis_eq_0: "cart_basis i = (0::'a::semiring_1^'n) \<longleftrightarrow> False"
- by (auto simp add: vec_eq_iff)
-
-lemma basis_nonzero: "cart_basis k \<noteq> (0:: 'a::semiring_1 ^'n)"
- by (simp add: basis_eq_0)
-
-text {* some lemmas to map between Eucl and Cart *}
-lemma basis_real_n[simp]:"(basis (\<pi>' i)::real^'a) = cart_basis i"
- unfolding basis_vec_def using pi'_range[where 'n='a]
- by (auto simp: vec_eq_iff axis_def)
-
-subsection {* Orthogonality on cartesian products *}
-
-lemma orthogonal_basis: "orthogonal (cart_basis i) x \<longleftrightarrow> x$i = (0::real)"
- by (auto simp add: orthogonal_def inner_vec_def cart_basis_def if_distrib
- cond_application_beta setsum_delta cong del: if_weak_cong)
-
-lemma orthogonal_basis_basis: "orthogonal (cart_basis i :: real^'n) (cart_basis j) \<longleftrightarrow> i \<noteq> j"
- unfolding orthogonal_basis[of i] basis_component[of j] by simp
-
-subsection {* Linearity on cartesian products *}
-
-lemma linear_vmul_component:
- assumes "linear f"
- shows "linear (\<lambda>x. f x $ k *\<^sub>R v)"
- using assms by (auto simp add: linear_def algebra_simps)
-
-
-subsection {* Adjoints on cartesian products *}
-
-text {* TODO: The following lemmas about adjoints should hold for any
-Hilbert space (i.e. complete inner product space).
-(see \url{http://en.wikipedia.org/wiki/Hermitian_adjoint})
-*}
-
-lemma adjoint_works_lemma:
- fixes f:: "real ^'n \<Rightarrow> real ^'m"
- assumes lf: "linear f"
- shows "\<forall>x y. f x \<bullet> y = x \<bullet> adjoint f y"
-proof -
- let ?N = "UNIV :: 'n set"
- let ?M = "UNIV :: 'm set"
- have fN: "finite ?N" by simp
- have fM: "finite ?M" by simp
- { fix y:: "real ^ 'm"
- let ?w = "(\<chi> i. (f (cart_basis i) \<bullet> y)) :: real ^ 'n"
- { fix x
- have "f x \<bullet> y = f (setsum (\<lambda>i. (x$i) *\<^sub>R cart_basis i) ?N) \<bullet> y"
- by (simp only: basis_expansion')
- also have "\<dots> = (setsum (\<lambda>i. (x$i) *\<^sub>R f (cart_basis i)) ?N) \<bullet> y"
- unfolding linear_setsum[OF lf fN]
- by (simp add: linear_cmul[OF lf])
- finally have "f x \<bullet> y = x \<bullet> ?w"
- by (simp add: inner_vec_def setsum_left_distrib
- setsum_right_distrib setsum_commute[of _ ?M ?N] field_simps)
- }
- }
- then show ?thesis
- unfolding adjoint_def some_eq_ex[of "\<lambda>f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y"]
- using choice_iff[of "\<lambda>a b. \<forall>x. f x \<bullet> a = x \<bullet> b "]
- by metis
-qed
-
-lemma adjoint_works:
- fixes f:: "real ^'n \<Rightarrow> real ^'m"
- assumes lf: "linear f"
- shows "x \<bullet> adjoint f y = f x \<bullet> y"
- using adjoint_works_lemma[OF lf] by metis
-
-lemma adjoint_linear:
- fixes f:: "real ^'n \<Rightarrow> real ^'m"
- assumes lf: "linear f"
- shows "linear (adjoint f)"
- unfolding linear_def vector_eq_ldot[where 'a="real^'n", symmetric] apply safe
- unfolding inner_simps smult_conv_scaleR adjoint_works[OF lf] by auto
-
-lemma adjoint_clauses:
- fixes f:: "real ^'n \<Rightarrow> real ^'m"
- assumes lf: "linear f"
- shows "x \<bullet> adjoint f y = f x \<bullet> y"
- and "adjoint f y \<bullet> x = y \<bullet> f x"
- by (simp_all add: adjoint_works[OF lf] inner_commute)
-
-lemma adjoint_adjoint:
- fixes f:: "real ^'n \<Rightarrow> real ^'m"
- assumes lf: "linear f"
- shows "adjoint (adjoint f) = f"
- by (rule adjoint_unique, simp add: adjoint_clauses [OF lf])
-
-
subsection {* Matrix operations *}
text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *}
@@ -680,10 +415,10 @@
apply auto
apply (subst vec_eq_iff)
apply clarify
- apply (clarsimp simp add: matrix_vector_mult_def cart_basis_def if_distrib cond_application_beta vec_eq_iff cong del: if_weak_cong)
- apply (erule_tac x="cart_basis ia" in allE)
+ apply (clarsimp simp add: matrix_vector_mult_def if_distrib cond_application_beta vec_eq_iff cong del: if_weak_cong)
+ apply (erule_tac x="axis ia 1" in allE)
apply (erule_tac x="i" in allE)
- apply (auto simp add: cart_basis_def if_distrib cond_application_beta
+ apply (auto simp add: if_distrib cond_application_beta axis_def
setsum_delta[OF finite] cong del: if_weak_cong)
done
@@ -728,25 +463,27 @@
by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult_commute)
lemma vector_componentwise:
- "(x::'a::ring_1^'n) = (\<chi> j. setsum (\<lambda>i. (x$i) * (cart_basis i :: 'a^'n)$j) (UNIV :: 'n set))"
- apply (subst basis_expansion[symmetric])
- apply (vector vec_eq_iff setsum_component)
- done
+ "(x::'a::ring_1^'n) = (\<chi> j. \<Sum>i\<in>UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)"
+ by (simp add: axis_def if_distrib setsum_cases vec_eq_iff)
+
+lemma basis_expansion: "setsum (\<lambda>i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)"
+ by (auto simp add: axis_def vec_eq_iff if_distrib setsum_cases cong del: if_weak_cong)
lemma linear_componentwise:
fixes f:: "real ^'m \<Rightarrow> real ^ _"
assumes lf: "linear f"
- shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (cart_basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
+ shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
proof -
let ?M = "(UNIV :: 'm set)"
let ?N = "(UNIV :: 'n set)"
have fM: "finite ?M" by simp
- have "?rhs = (setsum (\<lambda>i.(x$i) *\<^sub>R f (cart_basis i) ) ?M)$j"
- unfolding vector_smult_component[symmetric] smult_conv_scaleR
- unfolding setsum_component[of "(\<lambda>i.(x$i) *\<^sub>R f (cart_basis i :: real^'m))" ?M]
- ..
+ have "?rhs = (setsum (\<lambda>i.(x$i) *\<^sub>R f (axis i 1) ) ?M)$j"
+ unfolding setsum_component by simp
then show ?thesis
- unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion' ..
+ unfolding linear_setsum_mul[OF lf fM, symmetric]
+ unfolding scalar_mult_eq_scaleR[symmetric]
+ unfolding basis_expansion
+ by simp
qed
text{* Inverse matrices (not necessarily square) *}
@@ -761,7 +498,7 @@
text{* Correspondence between matrices and linear operators. *}
definition matrix :: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
- where "matrix f = (\<chi> i j. (f(cart_basis j))$i)"
+ where "matrix f = (\<chi> i j. (f(axis j 1))$i)"
lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::real ^ _))"
by (simp add: linear_def matrix_vector_mult_def vec_eq_iff
@@ -831,103 +568,8 @@
ultimately show ?thesis by metis
qed
-
-subsection {* Standard bases are a spanning set, and obviously finite. *}
-
-lemma span_stdbasis:"span {cart_basis i :: real^'n | i. i \<in> (UNIV :: 'n set)} = UNIV"
- apply (rule set_eqI)
- apply auto
- apply (subst basis_expansion'[symmetric])
- apply (rule span_setsum)
- apply simp
- apply auto
- apply (rule span_mul)
- apply (rule span_superset)
- apply auto
- done
-
-lemma finite_stdbasis: "finite {cart_basis i ::real^'n |i. i\<in> (UNIV:: 'n set)}" (is "finite ?S")
-proof -
- have "?S = cart_basis ` UNIV" by blast
- then show ?thesis by auto
-qed
-
-lemma card_stdbasis: "card {cart_basis i ::real^'n |i. i\<in> (UNIV :: 'n set)} = CARD('n)" (is "card ?S = _")
-proof -
- have "?S = cart_basis ` UNIV" by blast
- then show ?thesis using card_image[OF basis_inj] by simp
-qed
-
-lemma independent_stdbasis_lemma:
- assumes x: "(x::real ^ 'n) \<in> span (cart_basis ` S)"
- and iS: "i \<notin> S"
- shows "(x$i) = 0"
-proof -
- let ?U = "UNIV :: 'n set"
- let ?B = "cart_basis ` S"
- let ?P = "{(x::real^_). \<forall>i\<in> ?U. i \<notin> S \<longrightarrow> x$i =0}"
- { fix x::"real^_" assume xS: "x\<in> ?B"
- from xS have "x \<in> ?P" by auto }
- moreover
- have "subspace ?P"
- by (auto simp add: subspace_def)
- ultimately show ?thesis
- using x span_induct[of x ?B ?P] iS by blast
-qed
-
-lemma independent_stdbasis: "independent {cart_basis i ::real^'n |i. i\<in> (UNIV :: 'n set)}"
-proof -
- let ?I = "UNIV :: 'n set"
- let ?b = "cart_basis :: _ \<Rightarrow> real ^'n"
- let ?B = "?b ` ?I"
- have eq: "{?b i|i. i \<in> ?I} = ?B" by auto
- { assume d: "dependent ?B"
- then obtain k where k: "k \<in> ?I" "?b k \<in> span (?B - {?b k})"
- unfolding dependent_def by auto
- have eq1: "?B - {?b k} = ?B - ?b ` {k}" by simp
- have eq2: "?B - {?b k} = ?b ` (?I - {k})"
- unfolding eq1
- apply (rule inj_on_image_set_diff[symmetric])
- apply (rule basis_inj) using k(1)
- apply auto
- done
- from k(2) have th0: "?b k \<in> span (?b ` (?I - {k}))" unfolding eq2 .
- from independent_stdbasis_lemma[OF th0, of k, simplified]
- have False by simp }
- then show ?thesis unfolding eq dependent_def ..
-qed
-
lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *s b) = 0"
- unfolding inner_simps smult_conv_scaleR by auto
-
-lemma linear_eq_stdbasis_cart:
- assumes lf: "linear (f::real^'m \<Rightarrow> _)" and lg: "linear g"
- and fg: "\<forall>i. f (cart_basis i) = g(cart_basis i)"
- shows "f = g"
-proof -
- let ?U = "UNIV :: 'm set"
- let ?I = "{cart_basis i:: real^'m|i. i \<in> ?U}"
- { fix x assume x: "x \<in> (UNIV :: (real^'m) set)"
- from equalityD2[OF span_stdbasis]
- have IU: " (UNIV :: (real^'m) set) \<subseteq> span ?I" by blast
- from linear_eq[OF lf lg IU] fg x
- have "f x = g x" unfolding Ball_def mem_Collect_eq by metis
- }
- then show ?thesis by auto
-qed
-
-lemma bilinear_eq_stdbasis_cart:
- assumes bf: "bilinear (f:: real^'m \<Rightarrow> real^'n \<Rightarrow> _)"
- and bg: "bilinear g"
- and fg: "\<forall>i j. f (cart_basis i) (cart_basis j) = g (cart_basis i) (cart_basis j)"
- shows "f = g"
-proof -
- from fg have th: "\<forall>x \<in> {cart_basis i| i. i\<in> (UNIV :: 'm set)}.
- \<forall>y\<in> {cart_basis j |j. j \<in> (UNIV :: 'n set)}. f x y = g x y"
- by blast
- from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th]
- show ?thesis by blast
-qed
+ unfolding inner_simps scalar_mult_eq_scaleR by auto
lemma left_invertible_transpose:
"(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)"
@@ -1043,7 +685,7 @@
unfolding y[symmetric]
apply (rule span_setsum[OF fU])
apply clarify
- unfolding smult_conv_scaleR
+ unfolding scalar_mult_eq_scaleR
apply (rule span_mul)
apply (rule span_superset)
unfolding columns_def
@@ -1056,7 +698,7 @@
let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y"
{ fix y
have "?P y"
- proof (rule span_induct_alt[of ?P "columns A", folded smult_conv_scaleR])
+ proof (rule span_induct_alt[of ?P "columns A", folded scalar_mult_eq_scaleR])
show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
by (rule exI[where x=0], simp)
next
@@ -1159,25 +801,12 @@
dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc ..
-lemma infnorm_cart:"infnorm (x::real^'n) = Sup {abs(x$i) |i. i\<in> (UNIV :: 'n set)}"
- unfolding infnorm_def apply(rule arg_cong[where f=Sup]) apply safe
- apply(rule_tac x="\<pi> i" in exI) defer
- apply(rule_tac x="\<pi>' i" in exI)
- unfolding nth_conv_component
- using pi'_range apply auto
- done
-
-lemma infnorm_set_image_cart: "{abs(x$i) |i. i\<in> (UNIV :: _ set)} =
- (\<lambda>i. abs(x$i)) ` (UNIV)" by blast
-
-lemma infnorm_set_lemma_cart:
- "finite {abs((x::'a::abs ^'n)$i) |i. i\<in> (UNIV :: 'n set)}"
- "{abs(x$i) |i. i\<in> (UNIV :: 'n::finite set)} \<noteq> {}"
- unfolding infnorm_set_image_cart by auto
+lemma infnorm_cart:"infnorm (x::real^'n) = Sup {abs(x$i) |i. i\<in>UNIV}"
+ by (simp add: infnorm_def inner_axis Basis_vec_def) (metis (lifting) inner_axis real_inner_1_right)
lemma component_le_infnorm_cart: "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)"
- unfolding nth_conv_component
- using component_le_infnorm[of x] .
+ using Basis_le_infnorm[of "axis i 1" x]
+ by (simp add: Basis_vec_def axis_eq_axis inner_axis)
lemma continuous_component: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x $ i)"
unfolding continuous_def by (rule tendsto_vec_nth)
@@ -1371,7 +1000,7 @@
and "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i < c$i \<and> d$i < b$i)" (is ?th2)
and "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3)
and "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4)
- using subset_interval[of c d a b] by (simp_all add: cart_simps real_euclidean_nth)
+ using subset_interval[of c d a b] by (simp_all add: Basis_vec_def inner_axis)
lemma disjoint_interval_cart:
fixes a::"real^'n"
@@ -1379,7 +1008,7 @@
and "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2)
and "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i < c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th3)
and "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4)
- using disjoint_interval[of a b c d] by (simp_all add: cart_simps real_euclidean_nth)
+ using disjoint_interval[of a b c d] by (simp_all add: Basis_vec_def inner_axis)
lemma inter_interval_cart:
fixes a :: "'a::linorder^'n"
@@ -1400,7 +1029,7 @@
lemma is_interval_cart:
"is_interval (s::(real^'n) set) \<longleftrightarrow>
(\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i. ((a$i \<le> x$i \<and> x$i \<le> b$i) \<or> (b$i \<le> x$i \<and> x$i \<le> a$i))) \<longrightarrow> x \<in> s)"
- by (simp add: is_interval_def Ball_def cart_simps real_euclidean_nth)
+ by (simp add: is_interval_def Ball_def Basis_vec_def inner_axis imp_ex)
lemma closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \<le> a}"
by (simp add: closed_Collect_le)
@@ -1416,27 +1045,15 @@
lemma Lim_component_le_cart:
fixes f :: "'a \<Rightarrow> real^'n"
- assumes "(f ---> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. f(x)$i \<le> b) net"
+ assumes "(f ---> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. f x $i \<le> b) net"
shows "l$i \<le> b"
-proof -
- { fix x
- have "x \<in> {x::real^'n. inner (cart_basis i) x \<le> b} \<longleftrightarrow> x$i \<le> b"
- unfolding inner_basis by auto }
- then show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \<le> b}" f net l]
- using closed_halfspace_le[of "(cart_basis i)::real^'n" b] and assms(1,2,3) by auto
-qed
+ by (rule tendsto_le[OF assms(2) tendsto_const tendsto_vec_nth, OF assms(1, 3)])
lemma Lim_component_ge_cart:
fixes f :: "'a \<Rightarrow> real^'n"
assumes "(f ---> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. b \<le> (f x)$i) net"
shows "b \<le> l$i"
-proof -
- { fix x
- have "x \<in> {x::real^'n. inner (cart_basis i) x \<ge> b} \<longleftrightarrow> x$i \<ge> b"
- unfolding inner_basis by auto }
- then show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \<ge> b}" f net l]
- using closed_halfspace_ge[of b "(cart_basis i)::real^'n"] and assms(1,2,3) by auto
-qed
+ by (rule tendsto_le[OF assms(2) tendsto_vec_nth tendsto_const, OF assms(1, 3)])
lemma Lim_component_eq_cart:
fixes f :: "'a \<Rightarrow> real^'n"
@@ -1449,8 +1066,8 @@
lemma connected_ivt_component_cart:
fixes x :: "real^'n"
shows "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s. z$k = a)"
- using connected_ivt_hyperplane[of s x y "(cart_basis k)::real^'n" a]
- by (auto simp add: inner_basis)
+ using connected_ivt_hyperplane[of s x y "axis k 1" a]
+ by (auto simp add: inner_axis inner_commute)
lemma subspace_substandard_cart: "subspace {x::real^_. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
unfolding subspace_def by auto
@@ -1468,20 +1085,14 @@
lemma dim_substandard_cart: "dim {x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d"
(is "dim ?A = _")
proof -
- have *: "{x. \<forall>i<DIM((real, 'n) vec). i \<notin> \<pi>' ` d \<longrightarrow> x $$ i = 0} =
- {x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0}"
- apply safe
- apply (erule_tac x="\<pi>' i" in allE) defer
- apply (erule_tac x="\<pi> i" in allE)
- unfolding image_iff real_euclidean_nth[symmetric]
- apply (auto simp: pi'_inj[THEN inj_eq])
- done
- have " \<pi>' ` d \<subseteq> {..<DIM((real, 'n) vec)}"
- using pi'_range[where 'n='n] by auto
+ let ?a = "\<lambda>x. axis x 1 :: real^'n"
+ have *: "{x. \<forall>i\<in>Basis. i \<notin> ?a ` d \<longrightarrow> x \<bullet> i = 0} = ?A"
+ by (auto simp: image_iff Basis_vec_def axis_eq_axis inner_axis)
+ have "?a ` d \<subseteq> Basis"
+ by (auto simp: Basis_vec_def)
thus ?thesis
- using dim_substandard[of "\<pi>' ` d", where 'a="real^'n"]
- unfolding * using card_image[of "\<pi>'" d] using pi'_inj unfolding inj_on_def
- by auto
+ using dim_substandard[of "?a ` d"] card_image[of ?a d]
+ by (auto simp: axis_eq_axis inj_on_def *)
qed
lemma affinity_inverses:
@@ -1513,32 +1124,24 @@
using vector_affinity_eq[where m=m and x=x and y=y and c=c]
by metis
-lemma const_vector_cart:"((\<chi> i. d)::real^'n) = (\<chi>\<chi> i. d)"
- apply(subst euclidean_eq)
-proof safe
- case goal1
- hence *: "(basis i::real^'n) = cart_basis (\<pi> i)"
- unfolding basis_real_n[symmetric] by auto
- have "((\<chi> i. d)::real^'n) $$ i = d" unfolding euclidean_component_def *
- unfolding dot_basis by auto
- thus ?case using goal1 by auto
-qed
-
+lemma vector_cart:
+ fixes f :: "real^'n \<Rightarrow> real"
+ shows "(\<chi> i. f (axis i 1)) = (\<Sum>i\<in>Basis. f i *\<^sub>R i)"
+ unfolding euclidean_eq_iff[where 'a="real^'n"]
+ by simp (simp add: Basis_vec_def inner_axis)
+
+lemma const_vector_cart:"((\<chi> i. d)::real^'n) = (\<Sum>i\<in>Basis. d *\<^sub>R i)"
+ by (rule vector_cart)
subsection "Convex Euclidean Space"
-lemma Cart_1:"(1::real^'n) = (\<chi>\<chi> i. 1)"
- apply(subst euclidean_eq)
-proof safe
- case goal1
- thus ?case
- using nth_conv_component[THEN sym,where i1="\<pi> i" and x1="1::real^'n"] by auto
-qed
+lemma Cart_1:"(1::real^'n) = \<Sum>Basis"
+ using const_vector_cart[of 1] by (simp add: one_vec_def)
declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp]
declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp]
-lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta basis_component vector_uminus_component
+lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta vector_uminus_component
lemma convex_box_cart:
assumes "\<And>i. convex {x. P i x}"
@@ -1551,95 +1154,20 @@
lemma unit_interval_convex_hull_cart:
"{0::real^'n .. 1} = convex hull {x. \<forall>i. (x$i = 0) \<or> (x$i = 1)}" (is "?int = convex hull ?points")
unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"]
- apply(rule arg_cong[where f="\<lambda>x. convex hull x"]) apply(rule set_eqI) unfolding mem_Collect_eq
- apply safe apply(erule_tac x="\<pi>' i" in allE) unfolding nth_conv_component defer
- apply(erule_tac x="\<pi> i" in allE)
- apply auto
- done
+ by (rule arg_cong[where f="\<lambda>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 - (\<chi> i. d) .. x + (\<chi> i. d)} = convex hull s"
proof -
- obtain s where s: "finite s" "{x - (\<chi>\<chi> i. d)..x + (\<chi>\<chi> i. d)} = convex hull s"
- by (rule cube_convex_hull [OF assms])
- show thesis
- apply(rule that[OF s(1)]) unfolding s(2)[symmetric] const_vector_cart ..
+ from cube_convex_hull [OF assms, of x] guess s .
+ with that[of s] show thesis by (simp add: const_vector_cart)
qed
-lemma std_simplex_cart:
- "(insert (0::real^'n) { cart_basis i | i. i\<in>UNIV}) =
- (insert 0 { basis i | i. i<DIM((real,'n) vec)})"
- apply (rule arg_cong[where f="\<lambda>s. (insert 0 s)"])
- unfolding basis_real_n[symmetric]
- apply safe
- apply (rule_tac x="\<pi>' i" in exI) defer
- apply (rule_tac x="\<pi> i" in exI) using pi'_range[where 'n='n]
- apply auto
- done
-
-
-subsection "Brouwer Fixpoint"
-
-lemma kuhn_labelling_lemma_cart:
- assumes "(\<forall>x::real^_. P x \<longrightarrow> P (f x))" "\<forall>x. P x \<longrightarrow> (\<forall>i. Q i \<longrightarrow> 0 \<le> x$i \<and> x$i \<le> 1)"
- shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and>
- (\<forall>x i. P x \<and> Q i \<and> (x$i = 0) \<longrightarrow> (l x i = 0)) \<and>
- (\<forall>x i. P x \<and> Q i \<and> (x$i = 1) \<longrightarrow> (l x i = 1)) \<and>
- (\<forall>x i. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x$i \<le> f(x)$i) \<and>
- (\<forall>x i. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f(x)$i \<le> x$i)"
-proof -
- have and_forall_thm:"\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)"
- by auto
- have *: "\<forall>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 \<longrightarrow> (x \<noteq> 1 \<and> x \<le> y \<or> x \<noteq> 0 \<and> y \<le> x)"
- by auto
- show ?thesis
- unfolding and_forall_thm apply(subst choice_iff[symmetric])+
- proof (rule, rule)
- case goal1
- let ?R = "\<lambda>y. (P x \<and> Q xa \<and> x $ xa = 0 \<longrightarrow> y = (0::nat)) \<and>
- (P x \<and> Q xa \<and> x $ xa = 1 \<longrightarrow> y = 1) \<and>
- (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x $ xa \<le> f x $ xa) \<and>
- (P x \<and> Q xa \<and> y = 1 \<longrightarrow> f x $ xa \<le> x $ xa)"
- { assume "P x" "Q xa"
- hence "0 \<le> f x $ xa \<and> f x $ xa \<le> 1"
- using assms(2)[rule_format,of "f x" xa]
- apply (drule_tac assms(1)[rule_format])
- apply auto
- done
- }
- hence "?R 0 \<or> ?R 1" by auto
- thus ?case by auto
- qed
-qed
-
-lemma interval_bij_cart:"interval_bij = (\<lambda> (a,b) (u,v) (x::real^'n).
- (\<chi> i. u$i + (x$i - a$i) / (b$i - a$i) * (v$i - u$i))::real^'n)"
- unfolding interval_bij_def apply(rule ext)+ apply safe
- unfolding vec_eq_iff vec_lambda_beta unfolding nth_conv_component
- apply rule
- apply (subst euclidean_lambda_beta)
- using pi'_range apply auto
- done
-
-lemma interval_bij_affine_cart:
- "interval_bij (a,b) (u,v) = (\<lambda>x. (\<chi> i. (v$i - u$i) / (b$i - a$i) * x$i) +
- (\<chi> i. u$i - (v$i - u$i) / (b$i - a$i) * a$i)::real^'n)"
- apply rule
- unfolding vec_eq_iff interval_bij_cart vector_component_simps
- apply (auto simp add: field_simps add_divide_distrib[symmetric])
- done
-
subsection "Derivative"
-lemma has_derivative_vmul_component_cart:
- fixes c :: "real^'a \<Rightarrow> real^'b" and v :: "real^'c"
- assumes "(c has_derivative c') net"
- shows "((\<lambda>x. c(x)$k *\<^sub>R v) has_derivative (\<lambda>x. (c' x)$k *\<^sub>R v)) net"
- unfolding nth_conv_component by (intro has_derivative_intros assms)
-
lemma differentiable_at_imp_differentiable_on:
"(\<forall>x\<in>(s::(real^'n) set). f differentiable at x) \<Longrightarrow> f differentiable_on s"
unfolding differentiable_on_def by(auto intro!: differentiable_at_withinI)
@@ -1662,56 +1190,14 @@
subsection {* Component of the differential must be zero if it exists at a local
maximum or minimum for that corresponding component. *}
-lemma differential_zero_maxmin_component:
+lemma differential_zero_maxmin_cart:
fixes f::"real^'a \<Rightarrow> real^'b"
assumes "0 < e" "((\<forall>y \<in> ball x e. (f y)$k \<le> (f x)$k) \<or> (\<forall>y\<in>ball x e. (f x)$k \<le> (f y)$k))"
- "f differentiable (at x)" shows "jacobian f (at x) $ k = 0"
-(* FIXME: reuse proof of generic differential_zero_maxmin_component*)
-proof (rule ccontr)
- def D \<equiv> "jacobian f (at x)"
- assume "jacobian f (at x) $ k \<noteq> 0"
- then obtain j where j:"D$k$j \<noteq> 0" unfolding vec_eq_iff D_def by auto
- hence *: "abs (jacobian f (at x) $ k $ j) / 2 > 0"
- unfolding D_def by auto
- note as = assms(3)[unfolded jacobian_works has_derivative_at_alt]
- guess e' using as[THEN conjunct2,rule_format,OF *] .. note e' = this
- guess d using real_lbound_gt_zero[OF assms(1) e'[THEN conjunct1]] .. note d = this
- { fix c
- assume "abs c \<le> d"
- hence *:"norm (x + c *\<^sub>R cart_basis j - x) < e'"
- using norm_basis[of j] d by auto
- have "\<bar>(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\<bar> \<le>
- norm (f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j))"
- by (rule component_le_norm_cart)
- also have "\<dots> \<le> \<bar>D $ k $ j\<bar> / 2 * \<bar>c\<bar>"
- using e'[THEN conjunct2,rule_format,OF *] and norm_basis[of j]
- unfolding D_def[symmetric] by auto
- finally
- have "\<bar>(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\<bar> \<le>
- \<bar>D $ k $ j\<bar> / 2 * \<bar>c\<bar>" by simp
- hence "\<bar>f (x + c *\<^sub>R cart_basis j) $ k - f x $ k - c * D $ k $ j\<bar> \<le>
- \<bar>D $ k $ j\<bar> / 2 * \<bar>c\<bar>"
- unfolding vector_component_simps matrix_vector_mul_component
- unfolding smult_conv_scaleR[symmetric]
- unfolding inner_simps dot_basis smult_conv_scaleR by simp
- } note * = this
- have "x + d *\<^sub>R cart_basis j \<in> ball x e" "x - d *\<^sub>R cart_basis j \<in> ball x e"
- unfolding mem_ball dist_norm using norm_basis[of j] d by auto
- hence **: "((f (x - d *\<^sub>R cart_basis j))$k \<le> (f x)$k \<and> (f (x + d *\<^sub>R cart_basis j))$k \<le> (f x)$k) \<or>
- ((f (x - d *\<^sub>R cart_basis j))$k \<ge> (f x)$k \<and> (f (x + d *\<^sub>R cart_basis j))$k \<ge> (f x)$k)"
- using assms(2) by auto
- have ***: "\<And>y y1 y2 d dx::real. (y1\<le>y\<and>y2\<le>y) \<or> (y\<le>y1\<and>y\<le>y2) \<Longrightarrow>
- d < abs dx \<Longrightarrow> abs(y1 - y - - dx) \<le> d \<Longrightarrow> (abs (y2 - y - dx) \<le> d) \<Longrightarrow> False" by arith
- show False
- apply (rule ***[OF **, where dx="d * D $ k $ j" and d="\<bar>D $ k $ j\<bar> / 2 * \<bar>d\<bar>"])
- using *[of "-d"] and *[of d] and d[THEN conjunct1] and j
- unfolding mult_minus_left
- unfolding abs_mult diff_minus_eq_add scaleR_minus_left
- unfolding algebra_simps
- apply (auto intro: mult_pos_pos)
- done
-qed
-
+ "f differentiable (at x)"
+ shows "jacobian f (at x) $ k = 0"
+ using differential_zero_maxmin_component[of "axis k 1" e x f] assms
+ vector_cart[of "\<lambda>j. frechet_derivative f (at x) j $ k"]
+ by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def)
subsection {* Lemmas for working on @{typ "real^1"} *}
@@ -1775,25 +1261,6 @@
end
-(* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *)
-
-abbreviation vec1:: "'a \<Rightarrow> 'a ^ 1" where "vec1 x \<equiv> vec x"
-
-abbreviation dest_vec1:: "'a ^1 \<Rightarrow> 'a" where "dest_vec1 x \<equiv> (x$1)"
-
-lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x"
- by (simp add: vec_eq_iff)
-
-lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))"
- by (metis vec1_dest_vec1(1))
-
-lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))"
- by (metis vec1_dest_vec1(1))
-
-lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \<longleftrightarrow> x = y"
- by (metis vec1_dest_vec1(1))
-
-
subsection{* The collapse of the general concepts to dimension one. *}
lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
@@ -1863,438 +1330,10 @@
apply (simp add: forall_3)
done
-lemma range_vec1[simp]:"range vec1 = UNIV"
- apply (rule set_eqI,rule) unfolding image_iff defer
- apply (rule_tac x="dest_vec1 x" in bexI)
- apply auto
- done
-
-lemma dest_vec1_lambda: "dest_vec1(\<chi> i. x i) = x 1"
- by simp
-
-lemma dest_vec1_vec: "dest_vec1(vec x) = x"
- by simp
-
-lemma dest_vec1_sum: assumes fS: "finite S"
- shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S"
- apply (induct rule: finite_induct[OF fS])
- apply simp
- apply auto
- done
-
-lemma norm_vec1 [simp]: "norm(vec1 x) = abs(x)"
- by (simp add: vec_def norm_real)
-
-lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)"
- by (simp only: dist_real vec_component)
-lemma abs_dest_vec1: "norm x = \<bar>dest_vec1 x\<bar>"
- by (metis vec1_dest_vec1(1) norm_vec1)
-
-lemmas vec1_dest_vec1_simps =
- forall_vec1 vec_add[symmetric] dist_vec1 vec_sub[symmetric] vec1_dest_vec1 norm_vec1 vector_smult_component
- vec_inj[where 'b=1] vec_cmul[symmetric] smult_conv_scaleR[symmetric] o_def dist_real_def real_norm_def
-
-lemma bounded_linear_vec1: "bounded_linear (vec1::real\<Rightarrow>real^1)"
- unfolding bounded_linear_def additive_def bounded_linear_axioms_def
- unfolding smult_conv_scaleR[symmetric]
- unfolding vec1_dest_vec1_simps
- apply (rule conjI) defer
- apply (rule conjI) defer
- apply (rule_tac x=1 in exI)
- apply auto
- done
-
-lemma linear_vmul_dest_vec1:
- fixes f:: "real^_ \<Rightarrow> real^1"
- shows "linear f \<Longrightarrow> linear (\<lambda>x. dest_vec1(f x) *s v)"
- unfolding smult_conv_scaleR
- by (rule linear_vmul_component)
-
-lemma linear_from_scalars:
- assumes lf: "linear (f::real^1 \<Rightarrow> real^_)"
- shows "f = (\<lambda>x. dest_vec1 x *s column 1 (matrix f))"
- unfolding smult_conv_scaleR
- apply (rule ext)
- apply (subst matrix_works[OF lf, symmetric])
- apply (auto simp add: vec_eq_iff matrix_vector_mult_def column_def mult_commute)
- done
-
-lemma linear_to_scalars:
- assumes lf: "linear (f::real ^'n \<Rightarrow> real^1)"
- shows "f = (\<lambda>x. vec1(row 1 (matrix f) \<bullet> x))"
- apply (rule ext)
- apply (subst matrix_works[OF lf, symmetric])
- apply (simp add: vec_eq_iff matrix_vector_mult_def row_def inner_vec_def mult_commute)
- done
-
-lemma dest_vec1_eq_0: "dest_vec1 x = 0 \<longleftrightarrow> x = 0"
- by (simp add: dest_vec1_eq[symmetric])
-
-lemma setsum_scalars:
- assumes fS: "finite S"
- shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)"
- unfolding vec_setsum[OF fS] by simp
-
-lemma dest_vec1_wlog_le:
- "(\<And>(x::'a::linorder ^ 1) y. P x y \<longleftrightarrow> P y x)
- \<Longrightarrow> (\<And>x y. dest_vec1 x <= dest_vec1 y ==> P x y) \<Longrightarrow> P x y"
- apply (cases "dest_vec1 x \<le> dest_vec1 y")
- apply simp
- apply (subgoal_tac "dest_vec1 y \<le> dest_vec1 x")
- apply auto
- done
-
-text{* Lifting and dropping *}
-
-lemma continuous_on_o_dest_vec1:
- fixes f::"real \<Rightarrow> 'a::real_normed_vector"
- assumes "continuous_on {a..b::real} f"
- shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)"
- using assms unfolding continuous_on_iff apply safe
- apply (erule_tac x="x$1" in ballE,erule_tac x=e in allE)
- apply safe
- apply (rule_tac x=d in exI)
- apply safe
- unfolding o_def dist_real_def dist_real
- apply (erule_tac x="dest_vec1 x'" in ballE)
- apply (auto simp add:less_eq_vec_def)
- done
-
-lemma continuous_on_o_vec1:
- fixes f::"real^1 \<Rightarrow> 'a::real_normed_vector"
- assumes "continuous_on {a..b} f"
- shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)"
- using assms unfolding continuous_on_iff
- apply safe
- apply (erule_tac x="vec x" in ballE,erule_tac x=e in allE)
- apply safe
- apply (rule_tac x=d in exI)
- apply safe
- unfolding o_def dist_real_def dist_real
- apply (erule_tac x="vec1 x'" in ballE)
- apply (auto simp add:less_eq_vec_def)
- done
-
-lemma continuous_on_vec1:"continuous_on A (vec1::real\<Rightarrow>real^1)"
- by (rule linear_continuous_on[OF bounded_linear_vec1])
-
-lemma mem_interval_1:
- fixes x :: "real^1"
- shows "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
- and "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
- by (simp_all add: vec_eq_iff less_vec_def less_eq_vec_def)
-
-lemma vec1_interval:
- fixes a::"real"
- shows "vec1 ` {a .. b} = {vec1 a .. vec1 b}"
- and "vec1 ` {a<..<b} = {vec1 a<..<vec1 b}"
- apply (rule_tac[!] set_eqI)
- unfolding image_iff less_vec_def
- unfolding mem_interval_cart
- unfolding forall_1 vec1_dest_vec1_simps
- apply rule defer
- apply (rule_tac x="dest_vec1 x" in bexI) prefer 3
- apply rule defer
- apply (rule_tac x="dest_vec1 x" in bexI)
- apply auto
- done
-
-(* Some special cases for intervals in R^1. *)
-
-lemma interval_cases_1:
- fixes x :: "real^1"
- shows "x \<in> {a .. b} ==> x \<in> {a<..<b} \<or> (x = a) \<or> (x = b)"
- unfolding vec_eq_iff less_vec_def less_eq_vec_def mem_interval_cart
- by (auto simp del:dest_vec1_eq)
-
-lemma in_interval_1:
- fixes x :: "real^1"
- shows "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b) \<and>
- (x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
- unfolding vec_eq_iff less_vec_def less_eq_vec_def mem_interval_cart
- by (auto simp del:dest_vec1_eq)
-
-lemma interval_eq_empty_1:
- fixes a :: "real^1"
- shows "{a .. b} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a"
- and "{a<..<b} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
- unfolding interval_eq_empty_cart and ex_1 by auto
-
-lemma subset_interval_1:
- fixes a :: "real^1"
- shows "({a .. b} \<subseteq> {c .. d} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or>
- dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
- "({a .. b} \<subseteq> {c<..<d} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or>
- dest_vec1 c < dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b < dest_vec1 d)"
- "({a<..<b} \<subseteq> {c .. d} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or>
- dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
- "({a<..<b} \<subseteq> {c<..<d} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or>
- dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
- unfolding subset_interval_cart[of a b c d] unfolding forall_1 by auto
-
-lemma eq_interval_1:
- fixes a :: "real^1"
- shows "{a .. b} = {c .. d} \<longleftrightarrow>
- dest_vec1 b < dest_vec1 a \<and> dest_vec1 d < dest_vec1 c \<or>
- dest_vec1 a = dest_vec1 c \<and> dest_vec1 b = dest_vec1 d"
- unfolding set_eq_subset[of "{a .. b}" "{c .. d}"]
- unfolding subset_interval_1(1)[of a b c d]
- unfolding subset_interval_1(1)[of c d a b]
- by auto
-
-lemma disjoint_interval_1:
- fixes a :: "real^1"
- shows
- "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow>
- dest_vec1 b < dest_vec1 a \<or> dest_vec1 d < dest_vec1 c \<or> dest_vec1 b < dest_vec1 c \<or> dest_vec1 d < dest_vec1 a"
- "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow>
- dest_vec1 b < dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c \<or> dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
- "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow>
- dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d < dest_vec1 c \<or> dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
- "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow>
- dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c \<or> dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
- unfolding disjoint_interval_cart and ex_1 by auto
-
-lemma open_closed_interval_1:
- fixes a :: "real^1"
- shows "{a<..<b} = {a .. b} - {a, b}"
- unfolding set_eq_iff apply simp
- unfolding less_vec_def and less_eq_vec_def and forall_1 and dest_vec1_eq[symmetric]
- apply (auto simp del:dest_vec1_eq)
- done
-
-lemma closed_open_interval_1:
- "dest_vec1 (a::real^1) \<le> dest_vec1 b ==> {a .. b} = {a<..<b} \<union> {a,b}"
- unfolding set_eq_iff
- apply simp
- unfolding less_vec_def and less_eq_vec_def and forall_1 and dest_vec1_eq[symmetric]
- apply (auto simp del:dest_vec1_eq)
- done
-
-lemma Lim_drop_le:
- fixes f :: "'a \<Rightarrow> real^1"
- shows "(f ---> l) net \<Longrightarrow> \<not> trivial_limit net \<Longrightarrow>
- eventually (\<lambda>x. dest_vec1 (f x) \<le> b) net ==> dest_vec1 l \<le> b"
- using Lim_component_le_cart[of f l net 1 b] by auto
-
-lemma Lim_drop_ge:
- fixes f :: "'a \<Rightarrow> real^1"
- shows "(f ---> l) net \<Longrightarrow> \<not> trivial_limit net \<Longrightarrow>
- eventually (\<lambda>x. b \<le> dest_vec1 (f x)) net ==> b \<le> dest_vec1 l"
- using Lim_component_ge_cart[of f l net b 1] by auto
-
-
-text{* Also more convenient formulations of monotone convergence. *}
-
-lemma bounded_increasing_convergent:
- fixes s :: "nat \<Rightarrow> real^1"
- assumes "bounded {s n| n::nat. True}" "\<forall>n. dest_vec1(s n) \<le> dest_vec1(s(Suc n))"
- shows "\<exists>l. (s ---> l) sequentially"
-proof -
- obtain a where a:"\<forall>n. \<bar>dest_vec1 (s n)\<bar> \<le> a"
- using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto
- { fix m::nat
- have "\<And> n. n\<ge>m \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)"
- apply (induct_tac n)
- apply simp
- using assms(2) apply (erule_tac x="na" in allE)
- apply (auto simp add: not_less_eq_eq)
- done
- }
- hence "\<forall>m n. m \<le> n \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)"
- by auto
- then obtain l where "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>dest_vec1 (s n) - l\<bar> < e"
- using convergent_bounded_monotone[OF a] unfolding monoseq_def by auto
- thus ?thesis unfolding LIMSEQ_def apply(rule_tac x="vec1 l" in exI)
- unfolding dist_norm unfolding abs_dest_vec1 by auto
-qed
-
-lemma dest_vec1_simps[simp]:
- fixes a :: "real^1"
- shows "a$1 = 0 \<longleftrightarrow> a = 0" (*"a \<le> 1 \<longleftrightarrow> dest_vec1 a \<le> 1" "0 \<le> a \<longleftrightarrow> 0 \<le> dest_vec1 a"*)
- "a \<le> b \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 b" "dest_vec1 (1::real^1) = 1"
- by (auto simp add: less_eq_vec_def vec_eq_iff)
-
-lemma dest_vec1_inverval:
- "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}"
- "dest_vec1 ` {a<.. b} = {dest_vec1 a<.. dest_vec1 b}"
- "dest_vec1 ` {a ..<b} = {dest_vec1 a ..<dest_vec1 b}"
- "dest_vec1 ` {a<..<b} = {dest_vec1 a<..<dest_vec1 b}"
- apply(rule_tac [!] equalityI)
- unfolding subset_eq Ball_def Bex_def mem_interval_1 image_iff
- apply(rule_tac [!] allI)apply(rule_tac [!] impI)
- apply(rule_tac[2] x="vec1 x" in exI)apply(rule_tac[4] x="vec1 x" in exI)
- apply(rule_tac[6] x="vec1 x" in exI)apply(rule_tac[8] x="vec1 x" in exI)
- apply (auto simp add: less_vec_def less_eq_vec_def)
- done
-
-lemma dest_vec1_setsum:
- assumes "finite S"
- shows " dest_vec1 (setsum f S) = setsum (\<lambda>x. dest_vec1 (f x)) S"
- using dest_vec1_sum[OF assms] by auto
-
-lemma open_dest_vec1_vimage: "open S \<Longrightarrow> open (dest_vec1 -` S)"
- unfolding open_vec_def forall_1 by auto
-
-lemma tendsto_dest_vec1 [tendsto_intros]:
- "(f ---> l) net \<Longrightarrow> ((\<lambda>x. dest_vec1 (f x)) ---> dest_vec1 l) net"
- by (rule tendsto_vec_nth)
-
-lemma continuous_dest_vec1:
- "continuous net f \<Longrightarrow> continuous net (\<lambda>x. dest_vec1 (f x))"
- unfolding continuous_def by (rule tendsto_dest_vec1)
-
-lemma forall_dest_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P(dest_vec1 x))"
- apply safe defer
- apply (erule_tac x="vec1 x" in allE)
- apply auto
- done
-
-lemma forall_of_dest_vec1: "(\<forall>v. P (\<lambda>x. dest_vec1 (v x))) \<longleftrightarrow> (\<forall>x. P x)"
- apply rule
- apply rule
- apply (erule_tac x="vec1 \<circ> x" in allE)
- unfolding o_def vec1_dest_vec1
- apply auto
- done
-
-lemma forall_of_dest_vec1': "(\<forall>v. P (dest_vec1 v)) \<longleftrightarrow> (\<forall>x. P x)"
- apply rule
- apply rule
- apply (erule_tac x="(vec1 x)" in allE) defer
- apply rule
- apply (erule_tac x="dest_vec1 v" in allE)
- unfolding o_def vec1_dest_vec1
- apply auto
- done
-
-lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x"
- unfolding dist_norm by auto
-
-lemma bounded_linear_vec1_dest_vec1:
- fixes f :: "real \<Rightarrow> real"
- shows "linear (vec1 \<circ> f \<circ> dest_vec1) = bounded_linear f" (is "?l = ?r")
-proof -
- { assume ?l
- then have "\<exists>K. \<forall>x. norm ((vec1 \<circ> f \<circ> dest_vec1) x) \<le> K * norm x" by (rule linear_bounded)
- then guess K ..
- hence "\<exists>K. \<forall>x. \<bar>f x\<bar> \<le> \<bar>x\<bar> * K"
- apply(rule_tac x=K in exI)
- unfolding vec1_dest_vec1_simps by (auto simp add:field_simps)
- }
- thus ?thesis
- unfolding linear_def bounded_linear_def additive_def bounded_linear_axioms_def o_def
- unfolding vec1_dest_vec1_simps by auto
-qed
-
-lemma vec1_le[simp]: fixes a :: real shows "vec1 a \<le> vec1 b \<longleftrightarrow> a \<le> b"
- unfolding less_eq_vec_def by auto
-lemma vec1_less[simp]: fixes a :: real shows "vec1 a < vec1 b \<longleftrightarrow> a < b"
- unfolding less_vec_def by auto
-
-
-subsection {* Derivatives on real = Derivatives on @{typ "real^1"} *}
-
-lemma has_derivative_within_vec1_dest_vec1:
- fixes f :: "real \<Rightarrow> real"
- shows "((vec1 \<circ> f \<circ> dest_vec1) has_derivative (vec1 \<circ> f' \<circ> dest_vec1)) (at (vec1 x) within vec1 ` s)
- = (f has_derivative f') (at x within s)"
- unfolding has_derivative_within
- unfolding bounded_linear_vec1_dest_vec1[unfolded linear_conv_bounded_linear]
- unfolding o_def Lim_within Ball_def unfolding forall_vec1
- unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff
- by auto
-
-lemma has_derivative_at_vec1_dest_vec1:
- fixes f :: "real \<Rightarrow> real"
- shows "((vec1 \<circ> f \<circ> dest_vec1) has_derivative (vec1 \<circ> f' \<circ> dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)"
- using has_derivative_within_vec1_dest_vec1[where s=UNIV, unfolded range_vec1 within_UNIV]
- by auto
-
-lemma bounded_linear_vec1':
- fixes f :: "'a::real_normed_vector\<Rightarrow>real"
- shows "bounded_linear f = bounded_linear (vec1 \<circ> f)"
- unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def
- unfolding vec1_dest_vec1_simps by auto
-
-lemma bounded_linear_dest_vec1:
- fixes f :: "real\<Rightarrow>'a::real_normed_vector"
- shows "bounded_linear f = bounded_linear (f \<circ> dest_vec1)"
- unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def
- unfolding vec1_dest_vec1_simps
- by auto
-
-lemma has_derivative_at_vec1:
- fixes f :: "'a::real_normed_vector\<Rightarrow>real"
- shows "(f has_derivative f') (at x) = ((vec1 \<circ> f) has_derivative (vec1 \<circ> f')) (at x)"
- unfolding has_derivative_at
- unfolding bounded_linear_vec1'[unfolded linear_conv_bounded_linear]
- unfolding o_def Lim_at
- unfolding vec1_dest_vec1_simps dist_vec1_0
- by auto
-
-lemma has_derivative_within_dest_vec1:
- fixes f :: "real\<Rightarrow>'a::real_normed_vector"
- shows "((f \<circ> dest_vec1) has_derivative (f' \<circ> dest_vec1)) (at (vec1 x) within vec1 ` s) =
- (f has_derivative f') (at x within s)"
- unfolding has_derivative_within bounded_linear_dest_vec1
- unfolding o_def Lim_within Ball_def
- unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff
- by auto
-
-lemma has_derivative_at_dest_vec1:
- fixes f :: "real\<Rightarrow>'a::real_normed_vector"
- shows "((f \<circ> dest_vec1) has_derivative (f' \<circ> dest_vec1)) (at (vec1 x)) =
- (f has_derivative f') (at x)"
- using has_derivative_within_dest_vec1[where s=UNIV] by simp
-
-
-subsection {* In particular if we have a mapping into @{typ "real^1"}. *}
-
-lemma onorm_vec1:
- fixes f::"real \<Rightarrow> real"
- shows "onorm (\<lambda>x. vec1 (f (dest_vec1 x))) = onorm f"
-proof -
- have "\<forall>x::real^1. norm x = 1 \<longleftrightarrow> x\<in>{vec1 -1, vec1 (1::real)}"
- unfolding forall_vec1 by (auto simp add: vec_eq_iff)
- hence 1: "{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by auto
- have 2: "{norm (vec1 (f (dest_vec1 x))) |x. norm x = 1} =
- (\<lambda>x. norm (vec1 (f (dest_vec1 x)))) ` {x. norm x=1}"
- by auto
- have "\<forall>x::real. norm x = 1 \<longleftrightarrow> x\<in>{-1, 1}" by auto
- hence 3:"{x. norm x = 1} = {-1, (1::real)}" by auto
- have 4:"{norm (f x) |x. norm x = 1} = (\<lambda>x. norm (f x)) ` {x. norm x=1}" by auto
- show ?thesis
- unfolding onorm_def 1 2 3 4 by (simp add:Sup_finite_Max)
-qed
-
-lemma convex_vec1:"convex (vec1 ` s) = convex (s::real set)"
- unfolding convex_def Ball_def forall_vec1
- unfolding vec1_dest_vec1_simps image_iff
- by auto
-
lemma bounded_linear_component_cart[intro]: "bounded_linear (\<lambda>x::real^'n. x $ k)"
apply (rule bounded_linearI[where K=1])
using component_le_norm_cart[of _ k] unfolding real_norm_def by auto
-lemma bounded_vec1[intro]: "bounded s \<Longrightarrow> bounded (vec1 ` (s::real set))"
- unfolding bounded_def apply safe apply(rule_tac x="vec1 x" in exI,rule_tac x=e in exI)
- apply (auto simp add: dist_real dist_real_def)
- done
-
-(*lemma content_closed_interval_cases_cart:
- "content {a..b::real^'n} =
- (if {a..b} = {} then 0 else setprod (\<lambda>i. b$i - a$i) UNIV)"
-proof(cases "{a..b} = {}")
- case True thus ?thesis unfolding content_def by auto
-next case Falsethus ?thesis unfolding content_def unfolding if_not_P[OF False]
- proof(cases "\<forall>i. a $ i \<le> b $ i")
- case False thus ?thesis unfolding content_def using t by auto
- next case True note interval_eq_empty
- apply auto
-
- sorry*)
-
lemma integral_component_eq_cart[simp]:
fixes f :: "'n::ordered_euclidean_space \<Rightarrow> real^'m"
assumes "f integrable_on s"
@@ -2309,39 +1348,8 @@
unfolding vec_lambda_beta
by auto
-(*lemma content_split_cart:
- "content {a..b::real^'n} = content({a..b} \<inter> {x. x$k \<le> c}) + content({a..b} \<inter> {x. x$k >= c})"
-proof- note simps = interval_split_cart content_closed_interval_cases_cart vec_lambda_beta less_eq_vec_def
- { presume "a\<le>b \<Longrightarrow> ?thesis" thus ?thesis apply(cases "a\<le>b") unfolding simps by auto }
- have *:"UNIV = insert k (UNIV - {k})" "\<And>x. finite (UNIV-{x::'n})" "\<And>x. x\<notin>UNIV-{x}" by auto
- have *:"\<And>X Y Z. (\<Prod>i\<in>UNIV. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>UNIV-{k}. Z i (Y i))"
- "(\<Prod>i\<in>UNIV. b$i - a$i) = (\<Prod>i\<in>UNIV-{k}. b$i - a$i) * (b$k - a$k)"
- apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto
- assume as:"a\<le>b" moreover have "\<And>x. min (b $ k) c = max (a $ k) c
- \<Longrightarrow> x* (b$k - a$k) = x*(max (a $ k) c - a $ k) + x*(b $ k - max (a $ k) c)"
- by (auto simp add:field_simps)
- moreover have "\<not> a $ k \<le> c \<Longrightarrow> \<not> c \<le> b $ k \<Longrightarrow> False"
- unfolding not_le using as[unfolded less_eq_vec_def,rule_format,of k] by auto
- ultimately show ?thesis
- unfolding simps unfolding *(1)[of "\<lambda>i x. b$i - x"] *(1)[of "\<lambda>i x. x - a$i"] *(2) by(auto)
-qed*)
-
-lemma has_integral_vec1:
- assumes "(f has_integral k) {a..b}"
- shows "((\<lambda>x. vec1 (f x)) has_integral (vec1 k)) {a..b}"
-proof -
- have *: "\<And>p. (\<Sum>(x, k)\<in>p. content k *\<^sub>R vec1 (f x)) - vec1 k =
- vec1 ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k)"
- unfolding vec_sub vec_eq_iff by (auto simp add: split_beta)
- show ?thesis
- using assms unfolding has_integral
- apply safe
- apply(erule_tac x=e in allE,safe)
- apply(rule_tac x=d in exI,safe)
- apply(erule_tac x=p in allE,safe)
- unfolding * norm_vector_1
- apply auto
- done
-qed
+lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\<forall>i. a$i < b$i \<and> u$i < v$i"
+ shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x"
+ using assms by (intro interval_bij_bij) (auto simp: Basis_vec_def inner_axis)
end
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Dec 14 15:46:01 2012 +0100
@@ -101,56 +101,22 @@
lemma span_eq[simp]: "(span s = s) <-> subspace s"
unfolding span_def by (rule hull_eq, rule subspace_Inter)
-lemma basis_inj_on: "d \<subseteq> {..<DIM('n)} \<Longrightarrow> inj_on (basis :: nat => 'n::euclidean_space) d"
- by (auto simp add: inj_on_def euclidean_eq[where 'a='n])
-
-lemma finite_substdbasis: "finite {basis i ::'n::euclidean_space |i. i : (d:: nat set)}" (is "finite ?S")
-proof -
- have eq: "?S = basis ` d" by blast
- show ?thesis
- unfolding eq
- apply (rule finite_subset[OF _ range_basis_finite])
- apply auto
- done
-qed
-
-lemma card_substdbasis:
- assumes "d \<subseteq> {..<DIM('n::euclidean_space)}"
- shows "card {basis i ::'n::euclidean_space | i. i : d} = card d" (is "card ?S = _")
-proof -
- have eq: "?S = basis ` d" by blast
- show ?thesis
- unfolding eq
- using card_image[OF basis_inj_on[of d]] assms by auto
-qed
-
lemma substdbasis_expansion_unique:
- assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
- shows "setsum (%i. f i *\<^sub>R basis i) d = (x::'a::euclidean_space)
- <-> (!i<DIM('a). (i:d --> f i = x$$i) & (i ~: d --> x $$ i = 0))"
+ assumes d: "d \<subseteq> Basis"
+ shows "(\<Sum>i\<in>d. f i *\<^sub>R i) = (x::'a::euclidean_space)
+ \<longleftrightarrow> (\<forall>i\<in>Basis. (i \<in> d \<longrightarrow> f i = x \<bullet> i) \<and> (i \<notin> d \<longrightarrow> x \<bullet> i = 0))"
proof -
have *: "\<And>x a b P. x * (if P then a else b) = (if P then x*a else x*b)" by auto
- have **: "finite d" apply(rule finite_subset[OF assms]) by fastforce
- have ***: "\<And>i. (setsum (%i. f i *\<^sub>R ((basis i)::'a)) d) $$ i = (\<Sum>x\<in>d. if x = i then f x else 0)"
- unfolding euclidean_component_setsum euclidean_component_scaleR basis_component *
- apply (rule setsum_cong2)
- using assms apply auto
- done
+ have **: "finite d" by (auto intro: finite_subset[OF assms])
+ have ***: "\<And>i. i \<in> Basis \<Longrightarrow> (\<Sum>i\<in>d. f i *\<^sub>R i) \<bullet> i = (\<Sum>x\<in>d. if x = i then f x else 0)"
+ using d
+ by (auto intro!: setsum_cong simp: inner_Basis inner_setsum_left)
show ?thesis
- unfolding euclidean_eq[where 'a='a] *** setsum_delta[OF **] using assms by auto
-qed
-
-lemma independent_substdbasis:
- assumes "d \<subseteq> {..<DIM('a::euclidean_space)}"
- shows "independent {basis i ::'a::euclidean_space |i. i : (d :: nat set)}"
- (is "independent ?A")
-proof -
- have *: "{basis i |i. i < DIM('a)} = basis ` {..<DIM('a)}" by auto
- show ?thesis
- apply(intro independent_mono[of "{basis i ::'a |i. i : {..<DIM('a::euclidean_space)}}" "?A"] )
- using independent_basis[where 'a='a] assms apply (auto simp: *)
- done
-qed
+ unfolding euclidean_eq_iff[where 'a='a] by (auto simp: setsum_delta[OF **] ***)
+qed
+
+lemma independent_substdbasis: "d \<subseteq> Basis \<Longrightarrow> independent d"
+ by (rule independent_mono[OF independent_Basis])
lemma dim_cball:
assumes "0<e"
@@ -321,8 +287,8 @@
lemma vector_choose_size:
"0 <= c ==> \<exists>(x::'a::euclidean_space). norm x = c"
- apply (rule exI[where x="c *\<^sub>R basis 0 ::'a"])
- using DIM_positive[where 'a='a] apply auto
+ apply (rule exI[where x="c *\<^sub>R (SOME i. i \<in> Basis)"])
+ apply (auto simp: SOME_Basis)
done
lemma setsum_delta_notmem:
@@ -1291,11 +1257,12 @@
text {* Balls, being convex, are connected. *}
lemma convex_box: fixes a::"'a::euclidean_space"
- assumes "\<And>i. i<DIM('a) \<Longrightarrow> convex {x. P i x}"
- shows "convex {x. \<forall>i<DIM('a). P i (x$$i)}"
- using assms unfolding convex_def by auto
-
-lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i<DIM('a). 0 \<le> x$$i)}"
+ assumes "\<And>i. i\<in>Basis \<Longrightarrow> convex {x. P i x}"
+ shows "convex {x. \<forall>i\<in>Basis. P i (x\<bullet>i)}"
+ using assms unfolding convex_def
+ by (auto simp: inner_add_left)
+
+lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i)}"
by (rule convex_box) (simp add: atLeast_def[symmetric] convex_real_interval)
lemma convex_local_global_minimum:
@@ -2073,40 +2040,39 @@
from this show ?thesis using assms `span B=S` by auto
qed
-lemma span_substd_basis: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
- shows "(span {basis i | i. i : d}) = {x::'a::euclidean_space. (!i<DIM('a). i ~: d --> x$$i = 0)}"
- (is "span ?A = ?B")
+lemma span_substd_basis:
+ assumes d: "d \<subseteq> Basis"
+ shows "span d = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}" (is "_ = ?B")
proof-
-have "?A <= ?B" by auto
+have "d <= ?B" using d by (auto simp: inner_Basis)
moreover have s: "subspace ?B" using subspace_substandard[of "%i. i ~: d"] .
-ultimately have "span ?A <= ?B" using span_mono[of "?A" "?B"] span_eq[of "?B"] by blast
-moreover have "card d <= dim (span ?A)" using independent_card_le_dim[of "?A" "span ?A"]
- independent_substdbasis[OF assms] card_substdbasis[OF assms] span_inc[of "?A"] by auto
-moreover hence "dim ?B <= dim (span ?A)" using dim_substandard[OF assms] by auto
-ultimately show ?thesis using s subspace_dim_equal[of "span ?A" "?B"]
- subspace_span[of "?A"] by auto
+ultimately have "span d <= ?B" using span_mono[of d "?B"] span_eq[of "?B"] by blast
+moreover have "card d <= dim (span d)" using independent_card_le_dim[of d "span d"]
+ independent_substdbasis[OF assms] span_inc[of d] by auto
+moreover hence "dim ?B <= dim (span d)" using dim_substandard[OF assms] by auto
+ultimately show ?thesis using s subspace_dim_equal[of "span d" "?B"]
+ subspace_span[of d] by auto
qed
lemma basis_to_substdbasis_subspace_isomorphism:
fixes B :: "('a::euclidean_space) set"
assumes "independent B"
-shows "EX f d. card d = card B & linear f & f ` B = {basis i::'a |i. i : (d :: nat set)} &
- f ` span B = {x. ALL i<DIM('a). i ~: d --> x $$ i = (0::real)} & inj_on f (span B) \<and> d\<subseteq>{..<DIM('a)}"
+shows "EX f (d::'a set). card d = card B \<and> linear f \<and> f ` B = d \<and>
+ f ` span B = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} \<and> inj_on f (span B) \<and> d \<subseteq> Basis"
proof-
have B:"card B=dim B" using dim_unique[of B B "card B"] assms span_inc[of B] by auto
- def d \<equiv> "{..<dim B}" have t:"card d = dim B" unfolding d_def by auto
- have "dim B <= DIM('a)" using dim_subset_UNIV[of B] by auto
- hence d:"d\<subseteq>{..<DIM('a)}" unfolding d_def by auto
- let ?t = "{x::'a::euclidean_space. !i<DIM('a). i ~: d --> x$$i = 0}"
- have "EX f. linear f & f ` B = {basis i |i. i : d} &
- f ` span B = ?t & inj_on f (span B)"
- apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "{basis i |i. i : d}"])
+ have "dim B \<le> card (Basis :: 'a set)" using dim_subset_UNIV[of B] by simp
+ from ex_card[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B" by auto
+ let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i ~: d --> x\<bullet>i = 0}"
+ have "EX f. linear f & f ` B = d & f ` span B = ?t & inj_on f (span B)"
+ apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"])
apply(rule subspace_span) apply(rule subspace_substandard) defer
apply(rule span_inc) apply(rule assms) defer unfolding dim_span[of B] apply(rule B)
- unfolding span_substd_basis[OF d,symmetric] card_substdbasis[OF d] apply(rule span_inc)
+ unfolding span_substd_basis[OF d, symmetric]
+ apply(rule span_inc)
apply(rule independent_substdbasis[OF d]) apply(rule,assumption)
unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d] by auto
- from this t `card B=dim B` show ?thesis using d by auto
+ with t `card B = dim B` d show ?thesis by auto
qed
lemma aff_dim_empty:
@@ -2492,7 +2458,7 @@
using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] by (simp add: algebra_simps)
have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
unfolding dist_norm unfolding norm_scaleR[symmetric] apply(rule arg_cong[where f=norm]) using `e>0`
- by(auto simp add:euclidean_eq[where 'a='a] field_simps)
+ by(auto simp add:euclidean_eq_iff[where 'a='a] field_simps inner_simps)
also have "... = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
also have "... < d" using as[unfolded dist_norm] and `e>0`
by(auto simp add:pos_divide_less_eq[OF `e>0`] mult_commute)
@@ -2770,9 +2736,9 @@
subsection{* Some Properties of subset of standard basis *}
-lemma affine_hull_substd_basis: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
- shows "affine hull (insert 0 {basis i | i. i : d}) =
- {x::'a::euclidean_space. (!i<DIM('a). i ~: d --> x$$i = 0)}"
+lemma affine_hull_substd_basis: assumes "d\<subseteq>Basis"
+ shows "affine hull (insert 0 d) =
+ {x::'a::euclidean_space. (\<forall>i\<in>Basis. i ~: d --> x\<bullet>i = 0)}"
(is "affine hull (insert 0 ?A) = ?B")
proof- have *:"\<And>A. op + (0\<Colon>'a) ` A = A" "\<And>A. op + (- (0\<Colon>'a)) ` A = A" by auto
show ?thesis unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * ..
@@ -3230,10 +3196,10 @@
assumes "convex (s::('a::euclidean_space) set)" "closed s" "0 \<notin> s"
shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)"
proof(cases "s={}")
- case True have "norm ((basis 0)::'a) = 1" by auto
- hence "norm ((basis 0)::'a) = 1" "basis 0 \<noteq> (0::'a)" defer
- apply(subst norm_le_zero_iff[symmetric]) by auto
- thus ?thesis apply(rule_tac x="basis 0" in exI, rule_tac x=1 in exI)
+ case True
+ have "norm ((SOME i. i\<in>Basis)::'a) = 1" "(SOME i. i\<in>Basis) \<noteq> (0::'a)" defer
+ apply(subst norm_le_zero_iff[symmetric]) by (auto simp: SOME_Basis)
+ thus ?thesis apply(rule_tac x="SOME i. i\<in>Basis" in exI, rule_tac x=1 in exI)
using True using DIM_positive[where 'a='a] by auto
next case False thus ?thesis using False using separating_hyperplane_closed_point[OF assms]
apply - apply(erule exE)+ unfolding inner_zero_right apply(rule_tac x=a in exI, rule_tac x=b in exI) by auto qed
@@ -3703,10 +3669,10 @@
case False thus ?thesis apply (intro continuous_intros)
using cont_surfpi unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def by auto
next obtain B where B:"\<forall>x\<in>s. norm x \<le> B" using compact_imp_bounded[OF assms(1)] unfolding bounded_iff by auto
- hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis 0" in ballE) defer
- apply(erule_tac x="basis 0" in ballE)
+ hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="SOME i. i\<in>Basis" in ballE) defer
+ apply(erule_tac x="SOME i. i\<in>Basis" in ballE)
unfolding Ball_def mem_cball dist_norm using DIM_positive[where 'a='a]
- by auto
+ by (auto simp: SOME_Basis)
case True show ?thesis unfolding True continuous_at Lim_at apply(rule,rule) apply(rule_tac x="e / B" in exI)
apply(rule) apply(rule divide_pos_pos) prefer 3 apply(rule,rule,erule conjE)
unfolding norm_zero scaleR_zero_left dist_norm diff_0_right norm_scaleR abs_norm_cancel proof-
@@ -3849,7 +3815,8 @@
hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: field_simps)
hence "u * a + v * b \<le> b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) }
ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s" apply- apply(rule assms[unfolded is_interval_def, rule_format, OF as(1,2)])
- using as(3-) DIM_positive[where 'a='a] by auto qed
+ using as(3-) DIM_positive[where 'a='a] by (auto simp: inner_simps)
+qed
lemma is_interval_connected:
fixes s :: "('a::euclidean_space) set"
@@ -3892,8 +3859,8 @@
subsection {* Another intermediate value theorem formulation *}
lemma ivt_increasing_component_on_1: fixes f::"real \<Rightarrow> 'a::euclidean_space"
- assumes "a \<le> b" "continuous_on {a .. b} f" "(f a)$$k \<le> y" "y \<le> (f b)$$k"
- shows "\<exists>x\<in>{a..b}. (f x)$$k = y"
+ assumes "a \<le> b" "continuous_on {a .. b} f" "(f a)\<bullet>k \<le> y" "y \<le> (f b)\<bullet>k"
+ shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
proof- have "f a \<in> f ` {a..b}" "f b \<in> f ` {a..b}" apply(rule_tac[!] imageI)
using assms(1) by auto
thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y]
@@ -3902,20 +3869,20 @@
lemma ivt_increasing_component_1: fixes f::"real \<Rightarrow> 'a::euclidean_space"
shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a .. b}. continuous (at x) f
- \<Longrightarrow> f a$$k \<le> y \<Longrightarrow> y \<le> f b$$k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)$$k = y"
+ \<Longrightarrow> f a\<bullet>k \<le> y \<Longrightarrow> y \<le> f b\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>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 \<Rightarrow> 'a::euclidean_space"
- assumes "a \<le> b" "continuous_on {a .. b} f" "(f b)$$k \<le> y" "y \<le> (f a)$$k"
- shows "\<exists>x\<in>{a..b}. (f x)$$k = y"
+ assumes "a \<le> b" "continuous_on {a .. b} f" "(f b)\<bullet>k \<le> y" "y \<le> (f a)\<bullet>k"
+ shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
apply(subst neg_equal_iff_equal[symmetric])
using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"]
using assms using continuous_on_minus by auto
lemma ivt_decreasing_component_1: fixes f::"real \<Rightarrow> 'a::euclidean_space"
shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a .. b}. continuous (at x) f
- \<Longrightarrow> f b$$k \<le> y \<Longrightarrow> y \<le> f a$$k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)$$k = y"
+ \<Longrightarrow> f b\<bullet>k \<le> y \<Longrightarrow> y \<le> f a\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
by(rule ivt_decreasing_component_on_1)
(auto simp: continuous_at_imp_continuous_on)
@@ -3933,104 +3900,127 @@
thus "f x \<le> b" using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v]
unfolding obt(2-3) using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] by auto qed
+lemma inner_setsum_Basis[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1"
+ by (simp add: One_def inner_setsum_left setsum_cases inner_Basis)
+
lemma unit_interval_convex_hull:
- "{0::'a::ordered_euclidean_space .. (\<chi>\<chi> i. 1)} = convex hull {x. \<forall>i<DIM('a). (x$$i = 0) \<or> (x$$i = 1)}"
+ defines "One \<equiv> (\<Sum>Basis)"
+ shows "{0::'a::ordered_euclidean_space .. One} =
+ convex hull {x. \<forall>i\<in>Basis. (x\<bullet>i = 0) \<or> (x\<bullet>i = 1)}"
(is "?int = convex hull ?points")
-proof- have 01:"{0,(\<chi>\<chi> i. 1)} \<subseteq> convex hull ?points" apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto
- { fix n x assume "x\<in>{0::'a::ordered_euclidean_space .. \<chi>\<chi> i. 1}" "n \<le> DIM('a)" "card {i. i<DIM('a) \<and> x$$i \<noteq> 0} \<le> n"
+proof -
+ have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1"
+ by (simp add: One_def inner_setsum_left setsum_cases inner_Basis)
+ have 01:"{0,One} \<subseteq> convex hull ?points"
+ apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto
+ { fix n x assume "x\<in>{0::'a::ordered_euclidean_space .. One}" "n \<le> DIM('a)" "card {i. i\<in>Basis \<and> x\<bullet>i \<noteq> 0} \<le> n"
hence "x\<in>convex hull ?points" proof(induct n arbitrary: x)
- case 0 hence "x = 0" apply(subst euclidean_eq) apply rule by auto
+ case 0 hence "x = 0" apply(subst euclidean_eq_iff) apply rule by auto
thus "x\<in>convex hull ?points" using 01 by auto
next
- case (Suc n) show "x\<in>convex hull ?points" proof(cases "{i. i<DIM('a) \<and> x$$i \<noteq> 0} = {}")
- case True hence "x = 0" apply(subst euclidean_eq) by auto
+ case (Suc n) show "x\<in>convex hull ?points" proof(cases "{i. i\<in>Basis \<and> x\<bullet>i \<noteq> 0} = {}")
+ case True hence "x = 0" apply(subst euclidean_eq_iff) by auto
thus "x\<in>convex hull ?points" using 01 by auto
next
- case False def xi \<equiv> "Min ((\<lambda>i. x$$i) ` {i. i<DIM('a) \<and> x$$i \<noteq> 0})"
- have "xi \<in> (\<lambda>i. x$$i) ` {i. i<DIM('a) \<and> x$$i \<noteq> 0}" unfolding xi_def apply(rule Min_in) using False by auto
- then obtain i where i':"x$$i = xi" "x$$i \<noteq> 0" "i<DIM('a)" by auto
- have i:"\<And>j. j<DIM('a) \<Longrightarrow> x$$j > 0 \<Longrightarrow> x$$i \<le> x$$j"
+ case False def xi \<equiv> "Min ((\<lambda>i. x\<bullet>i) ` {i. i\<in>Basis \<and> x\<bullet>i \<noteq> 0})"
+ have "xi \<in> (\<lambda>i. x\<bullet>i) ` {i. i\<in>Basis \<and> x\<bullet>i \<noteq> 0}" unfolding xi_def apply(rule Min_in) using False by auto
+ then obtain i where i':"x\<bullet>i = xi" "x\<bullet>i \<noteq> 0" "i\<in>Basis" by auto
+ have i:"\<And>j. j\<in>Basis \<Longrightarrow> x\<bullet>j > 0 \<Longrightarrow> x\<bullet>i \<le> x\<bullet>j"
unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff
defer apply(rule_tac x=j in bexI) using i' by auto
- have i01:"x$$i \<le> 1" "x$$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i]
- using i'(2-) `x$$i \<noteq> 0` by auto
- show ?thesis proof(cases "x$$i=1")
- case True have "\<forall>j\<in>{i. i<DIM('a) \<and> x$$i \<noteq> 0}. x$$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq
- proof(erule conjE) fix j assume as:"x $$ j \<noteq> 0" "x $$ j \<noteq> 1" "j<DIM('a)"
- hence j:"x$$j \<in> {0<..<1}" using Suc(2) by(auto simp add: eucl_le[where 'a='a] elim!:allE[where x=j])
- hence "x$$j \<in> op $$ x ` {i. i<DIM('a) \<and> x $$ i \<noteq> 0}" using as(3) by auto
- hence "x$$j \<ge> x$$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto
+ have i01:"x\<bullet>i \<le> 1" "x\<bullet>i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i]
+ using i'(2-) `x\<bullet>i \<noteq> 0` by auto
+ show ?thesis proof(cases "x\<bullet>i=1")
+ case True have "\<forall>j\<in>{i. i\<in>Basis \<and> x\<bullet>i \<noteq> 0}. x\<bullet>j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq
+ proof(erule conjE) fix j assume as:"x \<bullet> j \<noteq> 0" "x \<bullet> j \<noteq> 1" "j\<in>Basis"
+ hence j:"x\<bullet>j \<in> {0<..<1}" using Suc(2)
+ by (auto simp add: eucl_le[where 'a='a] elim!:allE[where x=j])
+ hence "x\<bullet>j \<in> op \<bullet> x ` {i. i\<in>Basis \<and> x \<bullet> i \<noteq> 0}" using as(3) by auto
+ hence "x\<bullet>j \<ge> x\<bullet>i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto
thus False using True Suc(2) j by(auto simp add: elim!:ballE[where x=j]) qed
thus "x\<in>convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
by auto
- next let ?y = "\<lambda>j. if x$$j = 0 then 0 else (x$$j - x$$i) / (1 - x$$i)"
- case False hence *:"x = x$$i *\<^sub>R (\<chi>\<chi> j. if x$$j = 0 then 0 else 1) + (1 - x$$i) *\<^sub>R (\<chi>\<chi> j. ?y j)"
- apply(subst euclidean_eq) by(auto simp add: field_simps)
- { fix j assume j:"j<DIM('a)"
- have "x$$j \<noteq> 0 \<Longrightarrow> 0 \<le> (x $$ j - x $$ i) / (1 - x $$ i)" "(x $$ j - x $$ i) / (1 - x $$ i) \<le> 1"
+ next
+ let ?y = "\<Sum>j\<in>Basis. (if x\<bullet>j = 0 then 0 else (x\<bullet>j - x\<bullet>i) / (1 - x\<bullet>i)) *\<^sub>R j"
+ case False
+ then have *: "x = (x\<bullet>i) *\<^sub>R (\<Sum>j\<in>Basis. (if x\<bullet>j = 0 then 0 else 1) *\<^sub>R j) + (1 - x\<bullet>i) *\<^sub>R ?y"
+ by (subst euclidean_eq_iff) (simp add: inner_simps)
+ { fix j :: 'a assume j:"j\<in>Basis"
+ have "x\<bullet>j \<noteq> 0 \<Longrightarrow> 0 \<le> (x \<bullet> j - x \<bullet> i) / (1 - x \<bullet> i)" "(x \<bullet> j - x \<bullet> i) / (1 - x \<bullet> i) \<le> 1"
apply(rule_tac divide_nonneg_pos) using i(1)[of j] using False i01
using Suc(2)[unfolded mem_interval, rule_format, of j] using j
- by(auto simp add:field_simps)
- hence "0 \<le> ?y j \<and> ?y j \<le> 1" by auto }
- moreover have "i\<in>{j. j<DIM('a) \<and> x$$j \<noteq> 0} - {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0}"
+ by(auto simp add: field_simps)
+ with j have "0 \<le> ?y \<bullet> j \<and> ?y \<bullet> j \<le> 1" by (auto simp: inner_simps) }
+ moreover have "i\<in>{j. j\<in>Basis \<and> x\<bullet>j \<noteq> 0} - {j. j\<in>Basis \<and> ?y \<bullet> j \<noteq> 0}"
using i01 using i'(3) by auto
- hence "{j. j<DIM('a) \<and> x$$j \<noteq> 0} \<noteq> {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0}" using i'(3) by blast
- hence **:"{j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0} \<subset> {j. j<DIM('a) \<and> x$$j \<noteq> 0}" apply - apply rule
+ hence "{j. j\<in>Basis \<and> x\<bullet>j \<noteq> 0} \<noteq> {j. j\<in>Basis \<and> ?y \<bullet> j \<noteq> 0}" using i'(3) by blast
+ hence **:"{j. j\<in>Basis \<and> ?y \<bullet> j \<noteq> 0} \<subset> {j. j\<in>Basis \<and> x\<bullet>j \<noteq> 0}"
by auto
- have "card {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0} \<le> n"
+ have "card {j. j\<in>Basis \<and> ?y \<bullet> j \<noteq> 0} \<le> n"
using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto
- ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format])
- apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1))
- unfolding mem_interval using i01 Suc(3) by auto
- qed qed qed } note * = this
- have **:"DIM('a) = card {..<DIM('a)}" by auto
- show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule
- apply(rule_tac n2="DIM('a)" in *) prefer 3 apply(subst(2) **)
+ ultimately show ?thesis
+ apply(subst *)
+ apply(rule convex_convex_hull[unfolded convex_def, rule_format])
+ apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
+ defer
+ apply(rule Suc(1))
+ unfolding mem_interval
+ using i01 Suc(3)
+ by auto
+ qed
+ qed
+ qed } note * = this
+ show ?thesis
+ apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule
+ apply(rule_tac n2="DIM('a)" in *) prefer 3
apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule
- unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in allE)
- by auto qed
+ unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in ballE)
+ by auto
+qed
text {* And this is a finite set of vertices. *}
-lemma unit_cube_convex_hull: obtains s where "finite s" "{0 .. (\<chi>\<chi> i. 1)::'a::ordered_euclidean_space} = convex hull s"
- apply(rule that[of "{x::'a. \<forall>i<DIM('a). x$$i=0 \<or> x$$i=1}"])
- apply(rule finite_subset[of _ "(\<lambda>s. (\<chi>\<chi> i. if i\<in>s then 1::real else 0)::'a) ` Pow {..<DIM('a)}"])
+lemma unit_cube_convex_hull:
+ obtains s :: "'a::ordered_euclidean_space set" where "finite s" "{0 .. \<Sum>Basis} = convex hull s"
+ apply(rule that[of "{x::'a. \<forall>i\<in>Basis. x\<bullet>i=0 \<or> x\<bullet>i=1}"])
+ apply(rule finite_subset[of _ "(\<lambda>s. (\<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i)::'a) ` Pow Basis"])
prefer 3 apply(rule unit_interval_convex_hull) apply rule unfolding mem_Collect_eq proof-
- fix x::"'a" assume as:"\<forall>i<DIM('a). x $$ i = 0 \<or> x $$ i = 1"
- show "x \<in> (\<lambda>s. \<chi>\<chi> i. if i \<in> s then 1 else 0) ` Pow {..<DIM('a)}"
- apply(rule image_eqI[where x="{i. i<DIM('a) \<and> x$$i = 1}"])
- using as apply(subst euclidean_eq) by auto qed auto
+ fix x::"'a" assume as:"\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1"
+ show "x \<in> (\<lambda>s. \<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i) ` Pow Basis"
+ apply(rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
+ using as apply(subst euclidean_eq_iff) by (auto simp: inner_setsum_left_Basis)
+qed auto
text {* Hence any cube (could do any nonempty interval). *}
lemma cube_convex_hull:
assumes "0 < d" obtains s::"('a::ordered_euclidean_space) set" where
- "finite s" "{x - (\<chi>\<chi> i. d) .. x + (\<chi>\<chi> i. d)} = convex hull s" proof-
- let ?d = "(\<chi>\<chi> i. d)::'a"
- have *:"{x - ?d .. x + ?d} = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` {0 .. \<chi>\<chi> i. 1}" apply(rule set_eqI, rule)
+ "finite s" "{x - (\<Sum>i\<in>Basis. d*\<^sub>Ri) .. x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)} = convex hull s" proof-
+ let ?d = "(\<Sum>i\<in>Basis. d*\<^sub>Ri)::'a"
+ have *:"{x - ?d .. x + ?d} = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` {0 .. \<Sum>Basis}" apply(rule set_eqI, rule)
unfolding image_iff defer apply(erule bexE) proof-
fix y assume as:"y\<in>{x - ?d .. x + ?d}"
- { fix i assume i:"i<DIM('a)" have "x $$ i \<le> d + y $$ i" "y $$ i \<le> d + x $$ i"
- using as[unfolded mem_interval, THEN spec[where x=i]] i
- by auto
- hence "1 \<ge> inverse d * (x $$ i - y $$ i)" "1 \<ge> inverse d * (y $$ i - x $$ i)"
+ { fix i :: 'a assume i:"i\<in>Basis" have "x \<bullet> i \<le> d + y \<bullet> i" "y \<bullet> i \<le> d + x \<bullet> i"
+ using as[unfolded mem_interval, THEN bspec[where x=i]] i
+ by (auto simp: inner_simps)
+ hence "1 \<ge> inverse d * (x \<bullet> i - y \<bullet> i)" "1 \<ge> inverse d * (y \<bullet> i - x \<bullet> i)"
apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[symmetric]
using assms by(auto simp add: field_simps)
- hence "inverse d * (x $$ i * 2) \<le> 2 + inverse d * (y $$ i * 2)"
- "inverse d * (y $$ i * 2) \<le> 2 + inverse d * (x $$ i * 2)" by(auto simp add:field_simps) }
- hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> {0..\<chi>\<chi> i.1}" unfolding mem_interval using assms
- by(auto simp add: field_simps)
- thus "\<exists>z\<in>{0..\<chi>\<chi> i.1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI)
+ hence "inverse d * (x \<bullet> i * 2) \<le> 2 + inverse d * (y \<bullet> i * 2)"
+ "inverse d * (y \<bullet> i * 2) \<le> 2 + inverse d * (x \<bullet> i * 2)" by(auto simp add:field_simps) }
+ hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> {0..\<Sum>Basis}" unfolding mem_interval using assms
+ by(auto simp add: field_simps inner_simps)
+ thus "\<exists>z\<in>{0..\<Sum>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 by auto
next
- fix y z assume as:"z\<in>{0..\<chi>\<chi> i.1}" "y = x - ?d + (2*d) *\<^sub>R z"
- have "\<And>i. i<DIM('a) \<Longrightarrow> 0 \<le> d * z $$ i \<and> d * z $$ i \<le> d"
- using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE)
+ fix y z assume as:"z\<in>{0..\<Sum>Basis}" "y = x - ?d + (2*d) *\<^sub>R z"
+ have "\<And>i. i\<in>Basis \<Longrightarrow> 0 \<le> d * (z \<bullet> i) \<and> d * (z \<bullet> i) \<le> d"
+ using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in ballE)
apply rule apply(rule mult_nonneg_nonneg) prefer 3 apply(rule mult_right_le_one_le)
using assms by auto
thus "y \<in> {x - ?d..x + ?d}" unfolding as(2) mem_interval apply- apply rule using as(1)[unfolded mem_interval]
- apply(erule_tac x=i in allE) using assms by auto qed
- obtain s where "finite s" "{0::'a..\<chi>\<chi> i.1} = convex hull s" using unit_cube_convex_hull by auto
+ apply(erule_tac x=i in ballE) using assms by (auto simp: inner_simps) qed
+ obtain s where "finite s" "{0::'a..\<Sum>Basis} = convex hull s" using unit_cube_convex_hull by auto
thus ?thesis apply(rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"]) unfolding * and convex_hull_affinity by auto qed
subsection {* Bounded convex function on open set is continuous *}
@@ -4103,6 +4093,9 @@
subsubsection {* Hence a convex function on an open set is continuous *}
+lemma real_of_nat_ge_one_iff: "1 \<le> real (n::nat) \<longleftrightarrow> 1 \<le> n"
+ by auto
+
lemma convex_on_continuous:
assumes "open (s::('a::ordered_euclidean_space) set)" "convex_on s f"
(* FIXME: generalize to euclidean_space *)
@@ -4113,29 +4106,40 @@
then obtain e where e:"cball x e \<subseteq> s" "e>0" using assms(1) unfolding open_contains_cball by auto
def d \<equiv> "e / real DIM('a)"
have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto)
- let ?d = "(\<chi>\<chi> i. d)::'a"
+ let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a"
obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto
- have "x\<in>{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by auto
+ have "x\<in>{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by (auto simp: inner_setsum_left_Basis inner_simps)
hence "c\<noteq>{}" using c by auto
def k \<equiv> "Max (f ` c)"
- have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)])
- apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof
+ have "convex_on {x - ?d..x + ?d} f"
+ apply(rule convex_on_subset[OF assms(2)])
+ apply(rule subset_trans[OF _ e(1)])
+ unfolding subset_eq mem_cball
+ proof
fix z assume z:"z\<in>{x - ?d..x + ?d}"
- have e:"e = setsum (\<lambda>i. d) {..<DIM('a)}" unfolding setsum_constant d_def using dimge1
+ have e:"e = setsum (\<lambda>i::'a. d) Basis" unfolding setsum_constant d_def using dimge1
unfolding real_eq_of_nat by auto
show "dist x z \<le> e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono)
- using z[unfolded mem_interval] apply(erule_tac x=i in allE) by auto qed
+ using z[unfolded mem_interval] apply(erule_tac x=b in ballE) by (auto simp: inner_simps)
+ qed
hence k:"\<forall>y\<in>{x - ?d..x + ?d}. f y \<le> k" unfolding c(2) apply(rule_tac convex_on_convex_hull_bound) apply assumption
unfolding k_def apply(rule, rule Max_ge) using c(1) by auto
- have "d \<le> e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 by auto
+ have "d \<le> e"
+ unfolding d_def
+ apply(rule mult_imp_div_pos_le)
+ using `e>0`
+ unfolding mult_le_cancel_left1
+ apply (auto simp: real_of_nat_ge_one_iff Suc_le_eq DIM_positive)
+ done
hence dsube:"cball x d \<subseteq> cball x e" unfolding subset_eq Ball_def mem_cball by auto
have conv:"convex_on (cball x d) f" apply(rule convex_on_subset, rule convex_on_subset[OF assms(2)]) apply(rule e(1)) using dsube by auto
hence "\<forall>y\<in>cball x d. abs (f y) \<le> k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof
fix y assume y:"y\<in>cball x d"
- { fix i assume "i<DIM('a)" hence "x $$ i - d \<le> y $$ i" "y $$ i \<le> x $$ i + d"
- using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by auto }
+ { fix i :: 'a assume "i\<in>Basis" hence "x \<bullet> i - d \<le> y \<bullet> i" "y \<bullet> i \<le> x \<bullet> i + d"
+ using order_trans[OF Basis_le_norm y[unfolded mem_cball dist_norm], of i] by (auto simp: inner_diff_left) }
thus "f y \<le> k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm
- by auto qed
+ by (auto simp: inner_simps)
+ qed
hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous)
apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball)
apply force
@@ -4266,25 +4270,26 @@
have "norm (a - x) / norm (a - b) \<le> 1" unfolding divide_le_eq_1_pos[OF Fal2]
unfolding as[unfolded dist_norm] norm_ge_zero by auto
thus "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" apply(rule_tac x="dist a x / dist a b" in exI)
- unfolding dist_norm apply(subst euclidean_eq) apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4
- proof(rule,rule) fix i assume i:"i<DIM('a)"
- have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $$ i =
- ((norm (a - b) - norm (a - x)) * (a $$ i) + norm (a - x) * (b $$ i)) / norm (a - b)"
- using Fal by(auto simp add: field_simps)
- also have "\<dots> = x$$i" apply(rule divide_eq_imp[OF Fal])
+ unfolding dist_norm apply(subst euclidean_eq_iff) apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4
+ proof(rule) fix i :: 'a assume i:"i\<in>Basis"
+ have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i =
+ ((norm (a - b) - norm (a - x)) * (a \<bullet> i) + norm (a - x) * (b \<bullet> i)) / norm (a - b)"
+ using Fal by(auto simp add: field_simps inner_simps)
+ also have "\<dots> = x\<bullet>i" apply(rule divide_eq_imp[OF Fal])
unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq] apply-
- apply(subst (asm) euclidean_eq) using i apply(erule_tac x=i in allE) by(auto simp add:field_simps)
- finally show "x $$ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $$ i"
+ apply(subst (asm) euclidean_eq_iff) using i apply(erule_tac x=i in ballE) by(auto simp add:field_simps inner_simps)
+ finally show "x \<bullet> i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i"
by auto
- qed(insert Fal2, auto) qed qed
+ qed(insert Fal2, auto) qed
+qed
lemma between_midpoint: fixes a::"'a::euclidean_space" shows
"between (a,b) (midpoint a b)" (is ?t1)
"between (b,a) (midpoint a b)" (is ?t2)
proof- have *:"\<And>x y z. x = (1/2::real) *\<^sub>R z \<Longrightarrow> y = (1/2) *\<^sub>R z \<Longrightarrow> norm z = norm x + norm y" by auto
show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *)
- unfolding euclidean_eq[where 'a='a]
- by(auto simp add:field_simps) qed
+ unfolding euclidean_eq_iff[where 'a='a]
+ by(auto simp add:field_simps inner_simps) qed
lemma between_mem_convex_hull:
"between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}"
@@ -4303,7 +4308,7 @@
have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
unfolding dist_norm unfolding norm_scaleR[symmetric] apply(rule arg_cong[where f=norm]) using `e>0`
- by(auto simp add: euclidean_eq[where 'a='a] field_simps)
+ by(auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
also have "\<dots> = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
also have "\<dots> < d" using as[unfolded dist_norm] and `e>0`
by(auto simp add:pos_divide_less_eq[OF `e>0`] mult_commute)
@@ -4347,236 +4352,237 @@
apply(rule_tac x=u in exI) defer apply(rule_tac x="\<lambda>x. if x = 0 then 1 - setsum u s else u x" in exI) using assms(2)
unfolding if_smult and setsum_delta_notmem[OF assms(2)] by auto
-lemma substd_simplex: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
- shows "convex hull (insert 0 { basis i | i. i : d}) =
- {x::'a::euclidean_space . (!i<DIM('a). 0 <= x$$i) & setsum (%i. x$$i) d <= 1 &
- (!i<DIM('a). i ~: d --> x$$i = 0)}"
+lemma substd_simplex:
+ assumes d: "d \<subseteq> Basis"
+ shows "convex hull (insert 0 d) = {x. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d --> x\<bullet>i = 0)}"
(is "convex hull (insert 0 ?p) = ?s")
-(* Proof is a modified copy of the proof of similar lemma std_simplex in Convex_Euclidean_Space.thy *)
-proof- let ?D = d (*"{..<DIM('a::euclidean_space)}"*)
+proof- let ?D = d
have "0 ~: ?p" using assms by (auto simp: image_def)
- have "{(basis i)::'n::euclidean_space |i. i \<in> ?D} = basis ` ?D" by auto
- note sumbas = this setsum_reindex[OF basis_inj_on[of d], unfolded o_def, OF assms]
- show ?thesis unfolding simplex[OF finite_substdbasis `0 ~: ?p`]
+ from d have "finite d" by (blast intro: finite_subset finite_Basis)
+ show ?thesis unfolding simplex[OF `finite d` `0 ~: ?p`]
apply(rule set_eqI) unfolding mem_Collect_eq apply rule
apply(erule exE, (erule conjE)+) apply(erule_tac[2] conjE)+ proof-
- fix x::"'a::euclidean_space" and u assume as: "\<forall>x\<in>{basis i |i. i \<in>?D}. 0 \<le> u x"
- "setsum u {basis i |i. i \<in> ?D} \<le> 1" "(\<Sum>x\<in>{basis i |i. i \<in>?D}. u x *\<^sub>R x) = x"
- have *:"\<forall>i<DIM('a). i:d --> u (basis i) = x$$i" and "(!i<DIM('a). i ~: d --> x $$ i = 0)" using as(3)
- unfolding sumbas unfolding substdbasis_expansion_unique[OF assms] by auto
- hence **:"setsum u {basis i |i. i \<in> ?D} = setsum (op $$ x) ?D" unfolding sumbas
+ fix x::"'a::euclidean_space" and u assume as: "\<forall>x\<in>?D. 0 \<le> u x"
+ "setsum u ?D \<le> 1" "(\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
+ have *:"\<forall>i\<in>Basis. i:d --> u i = x\<bullet>i" and "(\<forall>i\<in>Basis. i ~: d --> x \<bullet> i = 0)" using as(3)
+ unfolding substdbasis_expansion_unique[OF assms] by auto
+ hence **:"setsum u ?D = setsum (op \<bullet> x) ?D"
apply-apply(rule setsum_cong2) using assms by auto
- have " (\<forall>i<DIM('a). 0 \<le> x$$i) \<and> setsum (op $$ x) ?D \<le> 1"
- apply - proof(rule,rule,rule)
- fix i assume i:"i<DIM('a)" have "i : d ==> 0 \<le> x$$i" unfolding *[rule_format,OF i,symmetric]
+ have " (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1"
+ apply - proof(rule,rule)
+ fix i :: 'a assume i:"i\<in>Basis" have "i : d ==> 0 \<le> x\<bullet>i" unfolding *[rule_format,OF i,symmetric]
apply(rule_tac as(1)[rule_format]) by auto
- moreover have "i ~: d ==> 0 \<le> x$$i"
- using `(!i<DIM('a). i ~: d --> x $$ i = 0)`[rule_format, OF i] by auto
- ultimately show "0 \<le> x$$i" by auto
+ moreover have "i ~: d ==> 0 \<le> x\<bullet>i"
+ using `(\<forall>i\<in>Basis. i ~: d --> x \<bullet> i = 0)`[rule_format, OF i] by auto
+ ultimately show "0 \<le> x\<bullet>i" by auto
qed(insert as(2)[unfolded **], auto)
- from this show " (\<forall>i<DIM('a). 0 \<le> x$$i) \<and> setsum (op $$ x) ?D \<le> 1 & (!i<DIM('a). i ~: d --> x $$ i = 0)"
- using `(!i<DIM('a). i ~: d --> x $$ i = 0)` by auto
- next fix x::"'a::euclidean_space" assume as:"\<forall>i<DIM('a). 0 \<le> x $$ i" "setsum (op $$ x) ?D \<le> 1"
- "(!i<DIM('a). i ~: d --> x $$ i = 0)"
- show "\<exists>u. (\<forall>x\<in>{basis i |i. i \<in> ?D}. 0 \<le> u x) \<and>
- setsum u {basis i |i. i \<in> ?D} \<le> 1 \<and> (\<Sum>x\<in>{basis i |i. i \<in> ?D}. u x *\<^sub>R x) = x"
- apply(rule_tac x="\<lambda>y. inner y x" in exI) apply(rule,rule) unfolding mem_Collect_eq apply(erule exE)
- using as(1) apply(erule_tac x=i in allE) unfolding sumbas apply safe unfolding not_less basis_zero
- unfolding substdbasis_expansion_unique[OF assms] euclidean_component_def[symmetric]
- using as(2,3) by(auto simp add:dot_basis not_less)
- qed qed
+ from this show " (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1 & (\<forall>i\<in>Basis. i ~: d --> x \<bullet> i = 0)"
+ using `(\<forall>i\<in>Basis. i ~: d --> x \<bullet> i = 0)` by auto
+ next fix x::"'a::euclidean_space" assume as:"\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "setsum (op \<bullet> x) ?D \<le> 1"
+ "(\<forall>i\<in>Basis. i ~: d --> x \<bullet> i = 0)"
+ show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> setsum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
+ using as d unfolding substdbasis_expansion_unique[OF assms]
+ by (rule_tac x="inner x" in exI) auto
+ qed
+qed
lemma std_simplex:
- "convex hull (insert 0 { basis i | i. i<DIM('a)}) =
- {x::'a::euclidean_space . (\<forall>i<DIM('a). 0 \<le> x$$i) \<and> setsum (\<lambda>i. x$$i) {..<DIM('a)} \<le> 1 }"
- using substd_simplex[of "{..<DIM('a)}"] by auto
+ "convex hull (insert 0 Basis) =
+ {x::'a::euclidean_space . (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis \<le> 1 }"
+ using substd_simplex[of Basis] by auto
lemma interior_std_simplex:
- "interior (convex hull (insert 0 { basis i| i. i<DIM('a)})) =
- {x::'a::euclidean_space. (\<forall>i<DIM('a). 0 < x$$i) \<and> setsum (\<lambda>i. x$$i) {..<DIM('a)} < 1 }"
+ "interior (convex hull (insert 0 Basis)) =
+ {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 < x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis < 1 }"
apply(rule set_eqI) unfolding mem_interior std_simplex unfolding subset_eq mem_Collect_eq Ball_def mem_ball
unfolding Ball_def[symmetric] apply rule apply(erule exE, (erule conjE)+) defer apply(erule conjE) proof-
- fix x::"'a" and e assume "0<e" and as:"\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x<DIM('a). 0 \<le> xa $$ x) \<and> setsum (op $$ xa) {..<DIM('a)} \<le> 1"
- show "(\<forall>xa<DIM('a). 0 < x $$ xa) \<and> setsum (op $$ x) {..<DIM('a)} < 1" apply(safe) proof-
- fix i assume i:"i<DIM('a)" thus "0 < x $$ i" using as[THEN spec[where x="x - (e / 2) *\<^sub>R basis i"]] and `e>0`
- unfolding dist_norm by (auto elim!:allE[where x=i])
- next have **:"dist x (x + (e / 2) *\<^sub>R basis 0) < e" using `e>0`
- unfolding dist_norm by(auto intro!: mult_strict_left_mono)
- have "\<And>i. i<DIM('a) \<Longrightarrow> (x + (e / 2) *\<^sub>R basis 0) $$ i = x$$i + (if i = 0 then e/2 else 0)"
- by auto
- hence *:"setsum (op $$ (x + (e / 2) *\<^sub>R basis 0)) {..<DIM('a)} = setsum (\<lambda>i. x$$i + (if 0 = i then e/2 else 0)) {..<DIM('a)}"
+ fix x::"'a" and e assume "0<e" and as:"\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x\<in>Basis. 0 \<le> xa \<bullet> x) \<and> setsum (op \<bullet> xa) Basis \<le> 1"
+ show "(\<forall>xa\<in>Basis. 0 < x \<bullet> xa) \<and> setsum (op \<bullet> x) Basis < 1" apply(safe) proof-
+ fix i :: 'a assume i:"i\<in>Basis" thus "0 < x \<bullet> i" using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and `e>0`
+ unfolding dist_norm
+ by (auto elim!:ballE[where x=i] simp: inner_simps)
+ next have **:"dist x (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) < e" using `e>0`
+ unfolding dist_norm by(auto intro!: mult_strict_left_mono simp: SOME_Basis)
+ have "\<And>i. i\<in>Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) \<bullet> i = x\<bullet>i + (if i = (SOME i. i\<in>Basis) then e/2 else 0)"
+ by (auto simp: SOME_Basis inner_Basis inner_simps)
+ hence *:"setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis = setsum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis"
apply(rule_tac setsum_cong) by auto
- have "setsum (op $$ x) {..<DIM('a)} < setsum (op $$ (x + (e / 2) *\<^sub>R basis 0)) {..<DIM('a)}" unfolding * setsum_addf
- using `0<e` DIM_positive[where 'a='a] apply(subst setsum_delta') by auto
+ have "setsum (op \<bullet> x) Basis < setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis" unfolding * setsum_addf
+ using `0<e` DIM_positive[where 'a='a] apply(subst setsum_delta') by (auto simp: SOME_Basis)
also have "\<dots> \<le> 1" using ** apply(drule_tac as[rule_format]) by auto
- finally show "setsum (op $$ x) {..<DIM('a)} < 1" by auto qed
-next fix x::"'a" assume as:"\<forall>i<DIM('a). 0 < x $$ i" "setsum (op $$ x) {..<DIM('a)} < 1"
+ finally show "setsum (op \<bullet> x) Basis < 1" by auto qed
+next fix x::"'a" assume as:"\<forall>i\<in>Basis. 0 < x \<bullet> i" "setsum (op \<bullet> x) Basis < 1"
guess a using UNIV_witness[where 'a='b] ..
- let ?d = "(1 - setsum (op $$ x) {..<DIM('a)}) / real (DIM('a))"
- have "Min ((op $$ x) ` {..<DIM('a)}) > 0" apply(rule Min_grI) using as(1) by auto
- moreover have"?d > 0" apply(rule divide_pos_pos) using as(2) by(auto simp add: Suc_le_eq)
- ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i<DIM('a). 0 \<le> y $$ i) \<and> setsum (op $$ y) {..<DIM('a)} \<le> 1"
- apply(rule_tac x="min (Min ((op $$ x) ` {..<DIM('a)})) ?D" in exI) apply rule defer apply(rule,rule) proof-
- fix y assume y:"dist x y < min (Min (op $$ x ` {..<DIM('a)})) ?d"
- have "setsum (op $$ y) {..<DIM('a)} \<le> setsum (\<lambda>i. x$$i + ?d) {..<DIM('a)}" proof(rule setsum_mono)
- fix i assume "i\<in>{..<DIM('a)}" hence "abs (y$$i - x$$i) < ?d" apply-apply(rule le_less_trans)
- using component_le_norm[of "y - x" i]
- using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add: norm_minus_commute)
- thus "y $$ i \<le> x $$ i + ?d" by auto qed
+ let ?d = "(1 - setsum (op \<bullet> x) Basis) / real (DIM('a))"
+ have "Min ((op \<bullet> x) ` Basis) > 0" apply(rule Min_grI) using as(1) by auto
+ moreover have"?d > 0" apply(rule divide_pos_pos) using as(2) by (auto simp add: Suc_le_eq DIM_positive)
+ ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
+ apply(rule_tac x="min (Min ((op \<bullet> x) ` Basis)) ?D" in exI) apply rule defer apply(rule,rule) proof-
+ fix y assume y:"dist x y < min (Min (op \<bullet> x ` Basis)) ?d"
+ have "setsum (op \<bullet> y) Basis \<le> setsum (\<lambda>i. x\<bullet>i + ?d) Basis" proof(rule setsum_mono)
+ fix i :: 'a assume i: "i\<in>Basis" hence "abs (y\<bullet>i - x\<bullet>i) < ?d" apply-apply(rule le_less_trans)
+ using Basis_le_norm[OF i, of "y - x"]
+ using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add: norm_minus_commute inner_diff_left)
+ thus "y \<bullet> i \<le> x \<bullet> i + ?d" by auto qed
also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant real_eq_of_nat by(auto simp add: Suc_le_eq)
- finally show "(\<forall>i<DIM('a). 0 \<le> y $$ i) \<and> setsum (op $$ y) {..<DIM('a)} \<le> 1"
- proof safe fix i assume i:"i<DIM('a)"
- have "norm (x - y) < x$$i" apply(rule less_le_trans)
+ finally show "(\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
+ proof safe fix i :: 'a assume i:"i\<in>Basis"
+ have "norm (x - y) < x\<bullet>i" apply(rule less_le_trans)
apply(rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]) using i by auto
- thus "0 \<le> y$$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by auto
+ thus "0 \<le> y\<bullet>i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i]
+ by (auto simp: inner_simps)
qed qed auto qed
lemma interior_std_simplex_nonempty: obtains a::"'a::euclidean_space" where
- "a \<in> interior(convex hull (insert 0 {basis i | i . i<DIM('a)}))" proof-
- let ?D = "{..<DIM('a)}" let ?a = "setsum (\<lambda>b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) {(basis i) | i. i<DIM('a)}"
- have *:"{basis i :: 'a | i. i<DIM('a)} = basis ` ?D" by auto
- { fix i assume i:"i<DIM('a)" have "?a $$ i = inverse (2 * real DIM('a))"
- unfolding euclidean_component_setsum * and setsum_reindex[OF basis_inj] and o_def
- apply(rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) apply(rule setsum_cong2)
- defer apply(subst setsum_delta') unfolding euclidean_component_def using i by(auto simp add:dot_basis) }
+ "a \<in> interior(convex hull (insert 0 Basis))" proof-
+ let ?D = "Basis :: 'a set" let ?a = "setsum (\<lambda>b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis"
+ { fix i :: 'a assume i:"i\<in>Basis" have "?a \<bullet> i = inverse (2 * real DIM('a))"
+ by (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"])
+ (simp_all add: setsum_cases i) }
note ** = this
show ?thesis apply(rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof safe
- fix i assume i:"i<DIM('a)" show "0 < ?a $$ i" unfolding **[OF i] by(auto simp add: Suc_le_eq)
- next have "setsum (op $$ ?a) ?D = setsum (\<lambda>i. inverse (2 * real DIM('a))) ?D" apply(rule setsum_cong2, rule **) by auto
+ fix i :: 'a assume i:"i\<in>Basis" show "0 < ?a \<bullet> i" unfolding **[OF i] by(auto simp add: Suc_le_eq DIM_positive)
+ next have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real DIM('a))) ?D" apply(rule setsum_cong2, rule **) by auto
also have "\<dots> < 1" unfolding setsum_constant real_eq_of_nat divide_inverse[symmetric] by (auto simp add:field_simps)
- finally show "setsum (op $$ ?a) ?D < 1" by auto qed qed
-
-lemma rel_interior_substd_simplex: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
- shows "rel_interior (convex hull (insert 0 { basis i| i. i : d})) =
- {x::'a::euclidean_space. (\<forall>i\<in>d. 0 < x$$i) & setsum (%i. x$$i) d < 1 & (!i<DIM('a). i ~: d --> x$$i = 0)}"
+ finally show "setsum (op \<bullet> ?a) ?D < 1" by auto qed qed
+
+lemma rel_interior_substd_simplex: assumes d: "d\<subseteq>Basis"
+ shows "rel_interior (convex hull (insert 0 d)) =
+ {x::'a::euclidean_space. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) < 1 \<and> (\<forall>i\<in>Basis. i ~: d --> x\<bullet>i = 0)}"
(is "rel_interior (convex hull (insert 0 ?p)) = ?s")
(* Proof is a modified copy of the proof of similar lemma interior_std_simplex in Convex_Euclidean_Space.thy *)
proof-
have "finite d" apply(rule finite_subset) using assms by auto
-{ assume "d={}" hence ?thesis using rel_interior_sing using euclidean_eq[of _ 0] by auto }
+{ assume "d={}" hence ?thesis using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto }
moreover
{ assume "d~={}"
-have h0: "affine hull (convex hull (insert 0 ?p))={x::'a::euclidean_space. (!i<DIM('a). i ~: d --> x$$i = 0)}"
+have h0: "affine hull (convex hull (insert 0 ?p))={x::'a::euclidean_space. (\<forall>i\<in>Basis. i ~: d --> x\<bullet>i = 0)}"
using affine_hull_convex_hull affine_hull_substd_basis assms by auto
-have aux: "!x::'n::euclidean_space. !i. ((! i:d. 0 <= x$$i) & (!i. i ~: d --> x$$i = 0))--> 0 <= x$$i" by auto
+have aux: "!!x::'a. \<forall>i\<in>Basis. ((\<forall>i\<in>d. 0 \<le> x\<bullet>i) \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)) \<longrightarrow> 0 \<le> x\<bullet>i"
+ by auto
{ fix x::"'a::euclidean_space" assume x_def: "x : rel_interior (convex hull (insert 0 ?p))"
from this obtain e where e0: "e>0" and
- "ball x e Int {xa. (!i<DIM('a). i ~: d --> xa$$i = 0)} <= convex hull (insert 0 ?p)"
+ "ball x e Int {xa. (\<forall>i\<in>Basis. i ~: d --> xa\<bullet>i = 0)} <= convex hull (insert 0 ?p)"
using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto
- hence as: "ALL xa. (dist x xa < e & (!i<DIM('a). i ~: d --> xa$$i = 0)) -->
- (!i : d. 0 <= xa $$ i) & setsum (op $$ xa) d <= 1"
+ hence as: "ALL xa. (dist x xa < e & (\<forall>i\<in>Basis. i ~: d --> xa\<bullet>i = 0)) -->
+ (!i : d. 0 <= xa \<bullet> i) & setsum (op \<bullet> xa) d <= 1"
unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto
- have x0: "(!i<DIM('a). i ~: d --> x$$i = 0)"
+ have x0: "(\<forall>i\<in>Basis. i ~: d --> x\<bullet>i = 0)"
using x_def rel_interior_subset substd_simplex[OF assms] by auto
- have "(!i : d. 0 < x $$ i) & setsum (op $$ x) d < 1 & (!i<DIM('a). i ~: d --> x$$i = 0)" apply(rule,rule)
+ have "(\<forall>i\<in>d. 0 < x \<bullet> i) & setsum (op \<bullet> x) d < 1 & (\<forall>i\<in>Basis. i ~: d --> x\<bullet>i = 0)" apply(rule,rule)
proof-
- fix i::nat assume "i:d"
- hence "\<forall>ia\<in>d. 0 \<le> (x - (e / 2) *\<^sub>R basis i) $$ ia" apply-apply(rule as[rule_format,THEN conjunct1])
- unfolding dist_norm using assms `e>0` x0 by auto
- thus "0 < x $$ i" apply(erule_tac x=i in ballE) using `e>0` `i\<in>d` assms by auto
+ fix i::'a assume "i\<in>d"
+ hence "\<forall>ia\<in>d. 0 \<le> (x - (e / 2) *\<^sub>R i) \<bullet> ia" apply-apply(rule as[rule_format,THEN conjunct1])
+ unfolding dist_norm using d `e>0` x0 by (auto simp: inner_simps inner_Basis)
+ thus "0 < x \<bullet> i" apply(erule_tac x=i in ballE) using `e>0` `i\<in>d` d
+ by (auto simp: inner_simps inner_Basis)
next obtain a where a:"a:d" using `d ~= {}` by auto
- have **:"dist x (x + (e / 2) *\<^sub>R basis a) < e"
- using `e>0` and Euclidean_Space.norm_basis[of a]
+ then have **:"dist x (x + (e / 2) *\<^sub>R a) < e"
+ using `e>0` norm_Basis[of a] d
unfolding dist_norm by auto
- have "\<And>i. (x + (e / 2) *\<^sub>R basis a) $$ i = x$$i + (if i = a then e/2 else 0)"
- unfolding euclidean_simps using a assms by auto
- hence *:"setsum (op $$ (x + (e / 2) *\<^sub>R basis a)) d =
- setsum (\<lambda>i. x$$i + (if a = i then e/2 else 0)) d" by(rule_tac setsum_cong, auto)
- have h1: "(ALL i<DIM('a). i ~: d --> (x + (e / 2) *\<^sub>R basis a) $$ i = 0)"
- using as[THEN spec[where x="x + (e / 2) *\<^sub>R basis a"]] `a:d` using x0
- by(auto elim:allE[where x=a])
- have "setsum (op $$ x) d < setsum (op $$ (x + (e / 2) *\<^sub>R basis a)) d" unfolding * setsum_addf
+ have "\<And>i. i\<in>Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = x\<bullet>i + (if i = a then e/2 else 0)"
+ using a d by (auto simp: inner_simps inner_Basis)
+ hence *:"setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d =
+ setsum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d" using d by (intro setsum_cong) auto
+ have "a \<in> Basis" using `a \<in> d` d by auto
+ then have h1: "(\<forall>i\<in>Basis. i ~: d --> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)"
+ using x0 d `a\<in>d` by (auto simp add: inner_add_left inner_Basis)
+ have "setsum (op \<bullet> x) d < setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d" unfolding * setsum_addf
using `0<e` `a:d` using `finite d` by(auto simp add: setsum_delta')
- also have "\<dots> \<le> 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R basis a"] by auto
- finally show "setsum (op $$ x) d < 1 & (!i<DIM('a). i ~: d --> x$$i = 0)" using x0 by auto
+ also have "\<dots> \<le> 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] by auto
+ finally show "setsum (op \<bullet> x) d < 1 & (\<forall>i\<in>Basis. i ~: d --> x\<bullet>i = 0)" using x0 by auto
qed
}
moreover
{
fix x::"'a::euclidean_space" assume as: "x : ?s"
- have "!i. ((0<x$$i) | (0=x$$i) --> 0<=x$$i)" by auto
+ have "!i. ((0<x\<bullet>i) | (0=x\<bullet>i) --> 0<=x\<bullet>i)" by auto
moreover have "!i. (i:d) | (i ~: d)" by auto
ultimately
- have "!i. ( (ALL i:d. 0 < x$$i) & (ALL i. i ~: d --> x$$i = 0) ) --> 0 <= x$$i" by metis
+ have "!i. ( (ALL i:d. 0 < x\<bullet>i) & (ALL i. i ~: d --> x\<bullet>i = 0) ) --> 0 <= x\<bullet>i" by metis
hence h2: "x : convex hull (insert 0 ?p)" using as assms
unfolding substd_simplex[OF assms] by fastforce
obtain a where a:"a:d" using `d ~= {}` by auto
- let ?d = "(1 - setsum (op $$ x) d) / real (card d)"
+ let ?d = "(1 - setsum (op \<bullet> x) d) / real (card d)"
have "0 < card d" using `d ~={}` `finite d` by (simp add: card_gt_0_iff)
- have "Min ((op $$ x) ` d) > 0" using as `d \<noteq> {}` `finite d` by (simp add: Min_grI)
+ have "Min ((op \<bullet> x) ` d) > 0" using as `d \<noteq> {}` `finite d` by (simp add: Min_grI)
moreover have "?d > 0" apply(rule divide_pos_pos) using as using `0 < card d` by auto
- ultimately have h3: "min (Min ((op $$ x) ` d)) ?d > 0" by auto
+ ultimately have h3: "min (Min ((op \<bullet> x) ` d)) ?d > 0" by auto
have "x : rel_interior (convex hull (insert 0 ?p))"
unfolding rel_interior_ball mem_Collect_eq h0 apply(rule,rule h2)
unfolding substd_simplex[OF assms]
- apply(rule_tac x="min (Min ((op $$ x) ` d)) ?d" in exI) apply(rule,rule h3) apply safe unfolding mem_ball
- proof- fix y::'a assume y:"dist x y < min (Min (op $$ x ` d)) ?d" and y2:"(!i<DIM('a). i ~: d --> y$$i = 0)"
- have "setsum (op $$ y) d \<le> setsum (\<lambda>i. x$$i + ?d) d" proof(rule setsum_mono)
- fix i assume i:"i\<in>d"
- have "abs (y$$i - x$$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i]
+ apply(rule_tac x="min (Min ((op \<bullet> x) ` d)) ?d" in exI) apply(rule,rule h3) apply safe unfolding mem_ball
+ proof-
+ fix y::'a assume y:"dist x y < min (Min (op \<bullet> x ` d)) ?d" and y2: "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> y\<bullet>i = 0"
+ have "setsum (op \<bullet> y) d \<le> setsum (\<lambda>i. x\<bullet>i + ?d) d"
+ proof(rule setsum_mono)
+ fix i assume "i \<in> d"
+ with d have i: "i \<in> Basis" by auto
+ have "abs (y\<bullet>i - x\<bullet>i) < ?d" apply(rule le_less_trans) using Basis_le_norm[OF i, of "y - x"]
using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2]
- by(auto simp add: norm_minus_commute)
- thus "y $$ i \<le> x $$ i + ?d" by auto qed
+ by (auto simp add: norm_minus_commute inner_simps)
+ thus "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
+ qed
also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant real_eq_of_nat
using `0 < card d` by auto
- finally show "setsum (op $$ y) d \<le> 1" .
-
- fix i assume "i<DIM('a)" thus "0 \<le> y$$i"
+ finally show "setsum (op \<bullet> y) d \<le> 1" .
+
+ fix i :: 'a assume i: "i \<in> Basis" thus "0 \<le> y\<bullet>i"
proof(cases "i\<in>d") case True
- have "norm (x - y) < x$$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
- using Min_gr_iff[of "op $$ x ` d" "norm (x - y)"] `0 < card d` `i:d`
+ have "norm (x - y) < x\<bullet>i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
+ using Min_gr_iff[of "op \<bullet> x ` d" "norm (x - y)"] `0 < card d` `i:d`
by (simp add: card_gt_0_iff)
- thus "0 \<le> y$$i" using component_le_norm[of "x - y" i] and as(1)[rule_format] by auto
+ thus "0 \<le> y\<bullet>i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format]
+ by (auto simp: inner_simps)
qed(insert y2, auto)
qed
} ultimately have
- "!!x :: 'a::euclidean_space. (x : rel_interior (convex hull insert 0 {basis i |i. i : d})) =
- (x : {x. (ALL i:d. 0 < x $$ i) &
- setsum (op $$ x) d < 1 & (ALL i<DIM('a). i ~: d --> x $$ i = 0)})" by blast
+ "\<And>x. (x : rel_interior (convex hull insert 0 d)) = (x \<in> {x. (ALL i:d. 0 < x \<bullet> i) &
+ setsum (op \<bullet> x) d < 1 & (\<forall>i\<in>Basis. i ~: d --> x \<bullet> i = 0)})" by blast
from this have ?thesis by (rule set_eqI)
} ultimately show ?thesis by blast
qed
-lemma rel_interior_substd_simplex_nonempty: assumes "d ~={}" "d\<subseteq>{..<DIM('a::euclidean_space)}"
+lemma rel_interior_substd_simplex_nonempty: assumes "d ~={}" "d\<subseteq>Basis"
obtains a::"'a::euclidean_space" where
- "a : rel_interior(convex hull (insert 0 {basis i | i . i : d}))" proof-
+ "a : rel_interior(convex hull (insert 0 d))" proof-
(* Proof is a modified copy of the proof of similar lemma interior_std_simplex_nonempty in Convex_Euclidean_Space.thy *)
- let ?D = d let ?a = "setsum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) {(basis i) | i. i \<in> ?D}"
- have *:"{basis i :: 'a | i. i \<in> ?D} = basis ` ?D" by auto
+ let ?D = d let ?a = "setsum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) ?D"
have "finite d" apply(rule finite_subset) using assms(2) by auto
hence d1: "0 < real(card d)" using `d ~={}` by auto
- { fix i assume "i:d" have "?a $$ i = inverse (2 * real (card d))"
- unfolding * setsum_reindex[OF basis_inj_on, OF assms(2)] o_def
+ { fix i assume "i:d"
+ have "?a \<bullet> i = inverse (2 * real (card d))"
apply(rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"])
- unfolding euclidean_component_setsum
+ unfolding inner_setsum_left
apply(rule setsum_cong2)
using `i:d` `finite d` setsum_delta'[of d i "(%k. inverse (2 * real (card d)))"] d1 assms(2)
- by (auto simp add: Euclidean_Space.basis_component[of i])}
+ by (auto simp: inner_simps inner_Basis set_rev_mp[OF _ assms(2)]) }
note ** = this
show ?thesis apply(rule that[of ?a]) unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq
proof safe fix i assume "i:d"
have "0 < inverse (2 * real (card d))" using d1 by auto
- also have "...=?a $$ i" using **[of i] `i:d` by auto
- finally show "0 < ?a $$ i" by auto
- next have "setsum (op $$ ?a) ?D = setsum (\<lambda>i. inverse (2 * real (card d))) ?D"
+ also have "...=?a \<bullet> i" using **[of i] `i:d` by auto
+ finally show "0 < ?a \<bullet> i" by auto
+ next have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real (card d))) ?D"
by(rule setsum_cong2, rule **)
also have "\<dots> < 1" unfolding setsum_constant real_eq_of_nat divide_real_def[symmetric]
by (auto simp add:field_simps)
- finally show "setsum (op $$ ?a) ?D < 1" by auto
- next fix i assume "i<DIM('a)" and "i~:d"
- have "?a : (span {basis i | i. i : d})"
- apply (rule span_setsum[of "{basis i |i. i : d}" "(%b. b /\<^sub>R (2 * real (card d)))" "{basis i |i. i : d}"])
- using finite_substdbasis[of d] apply blast
+ finally show "setsum (op \<bullet> ?a) ?D < 1" by auto
+ next fix i assume "i\<in>Basis" and "i~:d"
+ have "?a : (span d)"
+ apply (rule span_setsum[of d "(%b. b /\<^sub>R (2 * real (card d)))" d])
+ using finite_subset[OF assms(2) finite_Basis]
+ apply blast
proof-
- { fix x assume "(x :: 'a::euclidean_space): {basis i |i. i : d}"
- hence "x : span {basis i |i. i : d}"
- using span_superset[of _ "{basis i |i. i : d}"] by auto
- hence "(x /\<^sub>R (2 * real (card d))) : (span {basis i |i. i : d})"
- using span_mul[of x "{basis i |i. i : d}" "(inverse (real (card d)) / 2)"] by auto
- } thus "\<forall>x\<in>{basis i |i. i \<in> d}. x /\<^sub>R (2 * real (card d)) \<in> span {basis i ::'a |i. i \<in> d}" by auto
+ { fix x assume "(x :: 'a::euclidean_space): d"
+ hence "x : span d"
+ using span_superset[of _ "d"] by auto
+ hence "(x /\<^sub>R (2 * real (card d))) : (span d)"
+ using span_mul[of x "d" "(inverse (real (card d)) / 2)"] by auto
+ } thus "\<forall>x\<in>d. x /\<^sub>R (2 * real (card d)) \<in> span d" by auto
qed
- thus "?a $$ i = 0 " using `i~:d` unfolding span_substd_basis[OF assms(2)] using `i<DIM('a)` by auto
+ thus "?a \<bullet> i = 0 " using `i~:d` unfolding span_substd_basis[OF assms(2)] using `i\<in>Basis` by auto
qed
qed
@@ -4608,14 +4614,14 @@
ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S"
using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"]
assms hull_subset[of S] by auto
-obtain d and f::"'n=>'n" where fd: "card d = card B & linear f & f ` B = {basis i |i. i : (d :: nat set)} &
- f ` span B = {x. ALL i<DIM('n). i ~: d --> x $$ i = (0::real)} & inj_on f (span B)" and d:"d\<subseteq>{..<DIM('n)}"
+obtain d and f::"'n=>'n" where fd: "card d = card B & linear f & f ` B = d &
+ f ` span B = {x. \<forall>i\<in>Basis. i ~: d --> x \<bullet> i = (0::real)} & inj_on f (span B)" and d:"d\<subseteq>Basis"
using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B_def by auto
hence "bounded_linear f" using linear_conv_bounded_linear by auto
have "d ~={}" using fd B_def `B ~={}` by auto
-have "(insert 0 {basis i |i. i : d}) = f ` (insert 0 B)" using fd linear_0 by auto
-hence "(convex hull (insert 0 {basis i |i. i : d})) = f ` (convex hull (insert 0 B))"
- using convex_hull_linear_image[of f "(insert 0 {basis i |i. i : d})"]
+have "(insert 0 d) = f ` (insert 0 B)" using fd linear_0 by auto
+hence "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))"
+ using convex_hull_linear_image[of f "(insert 0 d)"]
convex_hull_linear_image[of f "(insert 0 B)"] `bounded_linear f` by auto
moreover have "rel_interior (f ` (convex hull insert 0 B)) =
f ` rel_interior (convex hull insert 0 B)"
--- a/src/HOL/Multivariate_Analysis/Derivative.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy Fri Dec 14 15:46:01 2012 +0100
@@ -158,9 +158,6 @@
lemmas mult_left_has_derivative =
bounded_linear.has_derivative [OF bounded_linear_mult_left]
-lemmas euclidean_component_has_derivative =
- bounded_linear.has_derivative [OF bounded_linear_euclidean_component]
-
lemma has_derivative_neg:
assumes "(f has_derivative f') net"
shows "((\<lambda>x. -(f x)) has_derivative (\<lambda>h. -(f' h))) net"
@@ -191,20 +188,12 @@
using assms by (induct, simp_all add: has_derivative_const has_derivative_add)
text {* Somewhat different results for derivative of scalar multiplier. *}
-(** move **)
-lemma linear_vmul_component: (* TODO: delete *)
- assumes lf: "linear f"
- shows "linear (\<lambda>x. f x $$ k *\<^sub>R v)"
- using lf
- by (auto simp add: linear_def algebra_simps)
-
lemmas has_derivative_intros =
has_derivative_id has_derivative_const
has_derivative_add has_derivative_sub has_derivative_neg
has_derivative_add_const
scaleR_left_has_derivative scaleR_right_has_derivative
inner_left_has_derivative inner_right_has_derivative
- euclidean_component_has_derivative
subsubsection {* Limit transformation for derivatives *}
@@ -531,7 +520,7 @@
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes "(f has_derivative f') (at x within s)"
assumes "(f has_derivative f'') (at x within s)"
- assumes "(\<forall>i<DIM('a). \<forall>e>0. \<exists>d. 0 < abs(d) \<and> abs(d) < e \<and> (x + d *\<^sub>R basis i) \<in> s)"
+ assumes "(\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < abs(d) \<and> abs(d) < e \<and> (x + d *\<^sub>R i) \<in> s)"
shows "f' = f''"
proof-
note as = assms(1,2)[unfolded has_derivative_def]
@@ -540,32 +529,32 @@
have "x islimpt s" unfolding islimpt_approachable
proof(rule,rule)
fix e::real assume "0<e" guess d
- using assms(3)[rule_format,OF DIM_positive `e>0`] ..
+ using assms(3)[rule_format,OF SOME_Basis `e>0`] ..
thus "\<exists>x'\<in>s. x' \<noteq> x \<and> dist x' x < e"
- apply(rule_tac x="x + d *\<^sub>R basis 0" in bexI)
- unfolding dist_norm by auto
+ apply(rule_tac x="x + d *\<^sub>R (SOME i. i \<in> Basis)" in bexI)
+ unfolding dist_norm by (auto simp: SOME_Basis nonzero_Basis)
qed
hence *:"netlimit (at x within s) = x" apply-apply(rule netlimit_within)
unfolding trivial_limit_within by simp
show ?thesis apply(rule linear_eq_stdbasis)
unfolding linear_conv_bounded_linear
apply(rule as(1,2)[THEN conjunct1])+
- proof(rule,rule,rule ccontr)
- fix i assume i:"i<DIM('a)" def e \<equiv> "norm (f' (basis i) - f'' (basis i))"
- assume "f' (basis i) \<noteq> f'' (basis i)"
+ proof(rule,rule ccontr)
+ fix i :: 'a assume i:"i \<in> Basis" def e \<equiv> "norm (f' i - f'' i)"
+ assume "f' i \<noteq> f'' i"
hence "e>0" unfolding e_def by auto
guess d using tendsto_diff [OF as(1,2)[THEN conjunct2], unfolded * Lim_within,rule_format,OF `e>0`] .. note d=this
guess c using assms(3)[rule_format,OF i d[THEN conjunct1]] .. note c=this
- have *:"norm (- ((1 / \<bar>c\<bar>) *\<^sub>R f' (c *\<^sub>R basis i)) + (1 / \<bar>c\<bar>) *\<^sub>R f'' (c *\<^sub>R basis i)) = norm ((1 / abs c) *\<^sub>R (- (f' (c *\<^sub>R basis i)) + f'' (c *\<^sub>R basis i)))"
+ have *:"norm (- ((1 / \<bar>c\<bar>) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \<bar>c\<bar>) *\<^sub>R f'' (c *\<^sub>R i)) = norm ((1 / abs c) *\<^sub>R (- (f' (c *\<^sub>R i)) + f'' (c *\<^sub>R i)))"
unfolding scaleR_right_distrib by auto
- also have "\<dots> = norm ((1 / abs c) *\<^sub>R (c *\<^sub>R (- (f' (basis i)) + f'' (basis i))))"
+ also have "\<dots> = norm ((1 / abs c) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))"
unfolding f'.scaleR f''.scaleR
unfolding scaleR_right_distrib scaleR_minus_right by auto
also have "\<dots> = e" unfolding e_def using c[THEN conjunct1]
- using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"]
+ using norm_minus_cancel[of "f' i - f'' i"]
by (auto simp add: add.commute ab_diff_minus)
finally show False using c
- using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R basis i"]
+ using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R i"]
unfolding dist_norm
unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff
scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib
@@ -584,37 +573,38 @@
lemma frechet_derivative_unique_within_closed_interval:
fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector"
- assumes "\<forall>i<DIM('a). a$$i < b$$i" "x \<in> {a..b}" (is "x\<in>?I")
+ assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" "x \<in> {a..b}" (is "x\<in>?I")
assumes "(f has_derivative f' ) (at x within {a..b})"
assumes "(f has_derivative f'') (at x within {a..b})"
shows "f' = f''"
apply(rule frechet_derivative_unique_within)
apply(rule assms(3,4))+
-proof(rule,rule,rule,rule)
- fix e::real and i assume "e>0" and i:"i<DIM('a)"
- thus "\<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R basis i \<in> {a..b}"
- proof(cases "x$$i=a$$i")
+proof(rule,rule,rule)
+ fix e::real and i :: 'a assume "e>0" and i:"i\<in>Basis"
+ thus "\<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R i \<in> {a..b}"
+ proof(cases "x\<bullet>i=a\<bullet>i")
case True thus ?thesis
- apply(rule_tac x="(min (b$$i - a$$i) e) / 2" in exI)
- using assms(1)[THEN spec[where x=i]] and `e>0` and assms(2)
- unfolding mem_interval euclidean_simps
- using i by (auto simp add: field_simps)
- next note * = assms(2)[unfolded mem_interval,THEN spec[where x=i]]
- case False moreover have "a $$ i < x $$ i" using False * by auto
+ apply(rule_tac x="(min (b\<bullet>i - a\<bullet>i) e) / 2" in exI)
+ using assms(1)[THEN bspec[where x=i]] and `e>0` and assms(2)
+ unfolding mem_interval
+ using i by (auto simp add: field_simps inner_simps inner_Basis)
+ next
+ note * = assms(2)[unfolded mem_interval, THEN bspec, OF i]
+ case False moreover have "a \<bullet> i < x \<bullet> i" using False * by auto
moreover {
- have "a $$ i * 2 + min (x $$ i - a $$ i) e \<le> a$$i *2 + x$$i - a$$i"
+ have "a \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e \<le> a\<bullet>i *2 + x\<bullet>i - a\<bullet>i"
by auto
- also have "\<dots> = a$$i + x$$i" by auto
- also have "\<dots> \<le> 2 * x$$i" using * by auto
- finally have "a $$ i * 2 + min (x $$ i - a $$ i) e \<le> x $$ i * 2" by auto
+ also have "\<dots> = a\<bullet>i + x\<bullet>i" by auto
+ also have "\<dots> \<le> 2 * (x\<bullet>i)" using * by auto
+ finally have "a \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e \<le> x \<bullet> i * 2" by auto
}
- moreover have "min (x $$ i - a $$ i) e \<ge> 0" using * and `e>0` by auto
- hence "x $$ i * 2 \<le> b $$ i * 2 + min (x $$ i - a $$ i) e" using * by auto
+ moreover have "min (x \<bullet> i - a \<bullet> i) e \<ge> 0" using * and `e>0` by auto
+ hence "x \<bullet> i * 2 \<le> b \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e" using * by auto
ultimately show ?thesis
- apply(rule_tac x="- (min (x$$i - a$$i) e) / 2" in exI)
- using assms(1)[THEN spec[where x=i]] and `e>0` and assms(2)
- unfolding mem_interval euclidean_simps
- using i by (auto simp add: field_simps)
+ apply(rule_tac x="- (min (x\<bullet>i - a\<bullet>i) e) / 2" in exI)
+ using assms(1)[THEN bspec, OF i] and `e>0` and assms(2)
+ unfolding mem_interval
+ using i by (auto simp add: field_simps inner_simps inner_Basis)
qed
qed
@@ -638,7 +628,7 @@
lemma frechet_derivative_within_closed_interval:
fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector"
- assumes "\<forall>i<DIM('a). a$$i < b$$i" and "x \<in> {a..b}"
+ assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" and "x \<in> {a..b}"
assumes "(f has_derivative f') (at x within {a.. b})"
shows "frechet_derivative f (at x within {a.. b}) = f'"
apply(rule frechet_derivative_unique_within_closed_interval[where f=f])
@@ -650,14 +640,14 @@
lemma linear_componentwise:
fixes f:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes lf: "linear f"
- shows "(f x) $$ j = (\<Sum>i<DIM('a). (x$$i) * (f (basis i)$$j))" (is "?lhs = ?rhs")
+ shows "(f x) \<bullet> j = (\<Sum>i\<in>Basis. (x\<bullet>i) * (f i\<bullet>j))" (is "?lhs = ?rhs")
proof -
- have fA: "finite {..<DIM('a)}" by simp
- have "?rhs = (\<Sum>i<DIM('a). x$$i *\<^sub>R f (basis i))$$j"
- by simp
+ have fA: "finite Basis" by simp
+ have "?rhs = (\<Sum>i\<in>Basis. (x\<bullet>i) *\<^sub>R (f i))\<bullet>j"
+ by (simp add: inner_setsum_left)
then show ?thesis
unfolding linear_setsum_mul[OF lf fA, symmetric]
- unfolding euclidean_representation[symmetric] ..
+ unfolding euclidean_representation ..
qed
text {* We do not introduce @{text jacobian}, which is defined on matrices, instead we use
@@ -665,52 +655,54 @@
lemma jacobian_works:
"(f::('a::euclidean_space) \<Rightarrow> ('b::euclidean_space)) differentiable net \<longleftrightarrow>
- (f has_derivative (\<lambda>h. \<chi>\<chi> i.
- \<Sum>j<DIM('a). frechet_derivative f net (basis j) $$ i * h $$ j)) net"
- (is "?differentiable \<longleftrightarrow> (f has_derivative (\<lambda>h. \<chi>\<chi> i. ?SUM h i)) net")
+ (f has_derivative (\<lambda>h. \<Sum>i\<in>Basis.
+ (\<Sum>j\<in>Basis. frechet_derivative f net (j) \<bullet> i * (h \<bullet> j)) *\<^sub>R i)) net"
+ (is "?differentiable \<longleftrightarrow> (f has_derivative (\<lambda>h. \<Sum>i\<in>Basis. ?SUM h i *\<^sub>R i)) net")
proof
assume *: ?differentiable
{ fix h i
- have "?SUM h i = frechet_derivative f net h $$ i" using *
+ have "?SUM h i = frechet_derivative f net h \<bullet> i" using *
by (auto intro!: setsum_cong
simp: linear_componentwise[of _ h i] linear_frechet_derivative) }
- thus "(f has_derivative (\<lambda>h. \<chi>\<chi> i. ?SUM h i)) net"
- using * by (simp add: frechet_derivative_works)
+ with * show "(f has_derivative (\<lambda>h. \<Sum>i\<in>Basis. ?SUM h i *\<^sub>R i)) net"
+ by (simp add: frechet_derivative_works euclidean_representation)
qed (auto intro!: differentiableI)
lemma differential_zero_maxmin_component:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
- assumes k: "k < DIM('b)"
- and ball: "0 < e" "((\<forall>y \<in> ball x e. (f y)$$k \<le> (f x)$$k) \<or> (\<forall>y\<in>ball x e. (f x)$$k \<le> (f y)$$k))"
+ assumes k: "k \<in> Basis"
+ and ball: "0 < e" "((\<forall>y \<in> ball x e. (f y)\<bullet>k \<le> (f x)\<bullet>k) \<or> (\<forall>y\<in>ball x e. (f x)\<bullet>k \<le> (f y)\<bullet>k))"
and diff: "f differentiable (at x)"
- shows "(\<chi>\<chi> j. frechet_derivative f (at x) (basis j) $$ k) = (0::'a)" (is "?D k = 0")
+ shows "(\<Sum>j\<in>Basis. (frechet_derivative f (at x) j \<bullet> k) *\<^sub>R j) = (0::'a)" (is "?D k = 0")
proof (rule ccontr)
assume "?D k \<noteq> 0"
- then obtain j where j: "?D k $$ j \<noteq> 0" "j < DIM('a)"
- unfolding euclidean_lambda_beta euclidean_eq[of _ "0::'a"] by auto
- hence *: "\<bar>?D k $$ j\<bar> / 2 > 0" by auto
+ then obtain j where j: "?D k \<bullet> j \<noteq> 0" "j \<in> Basis"
+ unfolding euclidean_eq_iff[of _ "0::'a"] by auto
+ hence *: "\<bar>?D k \<bullet> j\<bar> / 2 > 0" by auto
note as = diff[unfolded jacobian_works has_derivative_at_alt]
guess e' using as[THEN conjunct2, rule_format, OF *] .. note e' = this
guess d using real_lbound_gt_zero[OF ball(1) e'[THEN conjunct1]] .. note d = this
{ fix c assume "abs c \<le> d"
- hence *:"norm (x + c *\<^sub>R basis j - x) < e'" using norm_basis[of j] d by auto
- let ?v = "(\<chi>\<chi> i. \<Sum>l<DIM('a). ?D i $$ l * (c *\<^sub>R basis j :: 'a) $$ l)"
+ hence *:"norm (x + c *\<^sub>R j - x) < e'" using norm_Basis[OF j(2)] d by auto
+ let ?v = "(\<Sum>i\<in>Basis. (\<Sum>l\<in>Basis. ?D i \<bullet> l * ((c *\<^sub>R j :: 'a) \<bullet> l)) *\<^sub>R i)"
have if_dist: "\<And> P a b c. a * (if P then b else c) = (if P then a * b else a * c)" by auto
- have "\<bar>(f (x + c *\<^sub>R basis j) - f x - ?v) $$ k\<bar> \<le>
- norm (f (x + c *\<^sub>R basis j) - f x - ?v)" by (rule component_le_norm)
- also have "\<dots> \<le> \<bar>?D k $$ j\<bar> / 2 * \<bar>c\<bar>"
- using e'[THEN conjunct2, rule_format, OF *] and norm_basis[of j] by fastforce
- finally have "\<bar>(f (x + c *\<^sub>R basis j) - f x - ?v) $$ k\<bar> \<le> \<bar>?D k $$ j\<bar> / 2 * \<bar>c\<bar>" by simp
- hence "\<bar>f (x + c *\<^sub>R basis j) $$ k - f x $$ k - c * ?D k $$ j\<bar> \<le> \<bar>?D k $$ j\<bar> / 2 * \<bar>c\<bar>"
- unfolding euclidean_simps euclidean_lambda_beta using j k
- by (simp add: if_dist setsum_cases field_simps) } note * = this
- have "x + d *\<^sub>R basis j \<in> ball x e" "x - d *\<^sub>R basis j \<in> ball x e"
- unfolding mem_ball dist_norm using norm_basis[of j] d by auto
- hence **:"((f (x - d *\<^sub>R basis j))$$k \<le> (f x)$$k \<and> (f (x + d *\<^sub>R basis j))$$k \<le> (f x)$$k) \<or>
- ((f (x - d *\<^sub>R basis j))$$k \<ge> (f x)$$k \<and> (f (x + d *\<^sub>R basis j))$$k \<ge> (f x)$$k)" using ball by auto
+ have "\<bar>(f (x + c *\<^sub>R j) - f x - ?v) \<bullet> k\<bar> \<le>
+ norm (f (x + c *\<^sub>R j) - f x - ?v)" by (rule Basis_le_norm[OF k])
+ also have "\<dots> \<le> \<bar>?D k \<bullet> j\<bar> / 2 * \<bar>c\<bar>"
+ using e'[THEN conjunct2, rule_format, OF *] and norm_Basis[OF j(2)] j
+ by simp
+ finally have "\<bar>(f (x + c *\<^sub>R j) - f x - ?v) \<bullet> k\<bar> \<le> \<bar>?D k \<bullet> j\<bar> / 2 * \<bar>c\<bar>" by simp
+ hence "\<bar>f (x + c *\<^sub>R j) \<bullet> k - f x \<bullet> k - c * (?D k \<bullet> j)\<bar> \<le> \<bar>?D k \<bullet> j\<bar> / 2 * \<bar>c\<bar>"
+ using j k
+ by (simp add: inner_simps field_simps inner_Basis setsum_cases if_dist) }
+ note * = this
+ have "x + d *\<^sub>R j \<in> ball x e" "x - d *\<^sub>R j \<in> ball x e"
+ unfolding mem_ball dist_norm using norm_Basis[OF j(2)] d by auto
+ hence **:"((f (x - d *\<^sub>R j))\<bullet>k \<le> (f x)\<bullet>k \<and> (f (x + d *\<^sub>R j))\<bullet>k \<le> (f x)\<bullet>k) \<or>
+ ((f (x - d *\<^sub>R j))\<bullet>k \<ge> (f x)\<bullet>k \<and> (f (x + d *\<^sub>R j))\<bullet>k \<ge> (f x)\<bullet>k)" using ball by auto
have ***: "\<And>y y1 y2 d dx::real.
(y1\<le>y\<and>y2\<le>y) \<or> (y\<le>y1\<and>y\<le>y2) \<Longrightarrow> d < abs dx \<Longrightarrow> abs(y1 - y - - dx) \<le> d \<Longrightarrow> (abs (y2 - y - dx) \<le> d) \<Longrightarrow> False" by arith
- show False apply(rule ***[OF **, where dx="d * ?D k $$ j" and d="\<bar>?D k $$ j\<bar> / 2 * \<bar>d\<bar>"])
+ show False apply(rule ***[OF **, where dx="d * (?D k \<bullet> j)" and d="\<bar>?D k \<bullet> j\<bar> / 2 * \<bar>d\<bar>"])
using *[of "-d"] and *[of d] and d[THEN conjunct1] and j
unfolding mult_minus_left
unfolding abs_mult diff_minus_eq_add scaleR_minus_left
@@ -728,13 +720,13 @@
proof -
obtain e where e:"e>0" "ball x e \<subseteq> s"
using `open s`[unfolded open_contains_ball] and `x \<in> s` by auto
- with differential_zero_maxmin_component[where 'b=real, of 0 e x f, simplified]
- have "(\<chi>\<chi> j. frechet_derivative f (at x) (basis j)) = (0::'a)"
- unfolding differentiable_def using mono deriv by auto
+ with differential_zero_maxmin_component[where 'b=real, of 1 e x f] mono deriv
+ have "(\<Sum>j\<in>Basis. frechet_derivative f (at x) j *\<^sub>R j) = (0::'a)"
+ by (auto simp: Basis_real_def differentiable_def)
with frechet_derivative_at[OF deriv, symmetric]
- have "\<forall>i<DIM('a). f' (basis i) = 0"
- by (simp add: euclidean_eq[of _ "0::'a"])
- with derivative_is_linear[OF deriv, THEN linear_componentwise, of _ 0]
+ have "\<forall>i\<in>Basis. f' i = 0"
+ by (simp add: euclidean_eq_iff[of _ "0::'a"] inner_setsum_left_Basis)
+ with derivative_is_linear[OF deriv, THEN linear_componentwise, of _ 1]
show ?thesis by (simp add: fun_eq_iff)
qed
@@ -1281,8 +1273,8 @@
proof-
interpret bounded_linear g' using assms by auto
note f'g' = assms(4)[unfolded id_def o_def,THEN cong]
- have "g' (f' a (\<chi>\<chi> i.1)) = (\<chi>\<chi> i.1)" "(\<chi>\<chi> i.1) \<noteq> (0::'n)" defer
- apply(subst euclidean_eq) using f'g' by auto
+ have "g' (f' a (\<Sum>Basis)) = (\<Sum>Basis)" "(\<Sum>Basis) \<noteq> (0::'n)" defer
+ apply(subst euclidean_eq_iff) using f'g' by auto
hence *:"0 < onorm g'"
unfolding onorm_pos_lt[OF assms(3)[unfolded linear_linear]] by fastforce
def k \<equiv> "1 / onorm g' / 2" have *:"k>0" unfolding k_def using * by auto
@@ -1726,7 +1718,7 @@
have *:"(\<lambda>x. x *\<^sub>R f') = (\<lambda>x. x *\<^sub>R f'')"
apply(rule frechet_derivative_unique_within_closed_interval[of "a" "b"])
using assms(3-)[unfolded has_vector_derivative_def] using assms(1-2)
- by auto
+ by (auto simp: Basis_real_def)
show ?thesis
proof(rule ccontr)
assume "f' \<noteq> f''"
--- a/src/HOL/Multivariate_Analysis/Determinants.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Determinants.thy Fri Dec 14 15:46:01 2012 +0100
@@ -452,7 +452,7 @@
ultimately show ?thesis
apply -
- apply (rule span_induct_alt[of ?P ?S, OF P0, folded smult_conv_scaleR])
+ apply (rule span_induct_alt[of ?P ?S, OF P0, folded scalar_mult_eq_scaleR])
apply blast
apply (rule x)
done
@@ -746,7 +746,7 @@
apply (rule span_setsum)
apply simp
apply (rule ballI)
- apply (rule span_mul [where 'a="real^'n", folded smult_conv_scaleR])+
+ apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+
apply (rule span_superset)
apply auto
done
@@ -782,7 +782,7 @@
apply (rule det_row_span)
apply (rule span_setsum[OF fUk])
apply (rule ballI)
- apply (rule span_mul [where 'a="real^'n", folded smult_conv_scaleR])+
+ apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+
apply (rule span_superset)
apply auto
done
@@ -879,9 +879,10 @@
have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)"
"\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
by simp_all
- from fd[rule_format, of "cart_basis i" "cart_basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
+ from fd[rule_format, of "axis i 1" "axis j 1", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
have "?A$i$j = ?m1 $ i $ j"
- by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def cart_basis_def th0 setsum_delta[OF fU] mat_def)}
+ by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def
+ th0 setsum_delta[OF fU] mat_def axis_def) }
hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector
with lf have ?rhs by blast}
moreover
@@ -931,7 +932,9 @@
unfolding dot_norm_neg dist_norm[symmetric]
unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)}
note fc = this
- show ?thesis unfolding linear_def vector_eq[where 'a="real^'n"] smult_conv_scaleR by (simp add: inner_add fc field_simps)
+ show ?thesis
+ unfolding linear_def vector_eq[where 'a="real^'n"] scalar_mult_eq_scaleR
+ by (simp add: inner_add fc field_simps)
qed
lemma isometry_linear:
@@ -981,7 +984,7 @@
apply (subst H(4))
using H(5-9)
apply (simp add: norm_eq norm_eq_1)
- apply (simp add: inner_diff smult_conv_scaleR) unfolding *
+ apply (simp add: inner_diff scalar_mult_eq_scaleR) unfolding *
by (simp add: field_simps) }
note th0 = this
let ?g = "\<lambda>x. if x = 0 then 0 else norm x *\<^sub>R f (inverse (norm x) *\<^sub>R x)"
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Fri Dec 14 15:46:01 2012 +0100
@@ -23,24 +23,24 @@
assumes euclidean_all_zero_iff:
"(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> (x = 0)"
- -- "FIXME: make this a separate definition"
- fixes dimension :: "'a itself \<Rightarrow> nat"
- assumes dimension_def: "dimension TYPE('a) = card Basis"
-
- -- "FIXME: eventually basis function can be removed"
- fixes basis :: "nat \<Rightarrow> 'a"
- assumes image_basis: "basis ` {..<dimension TYPE('a)} = Basis"
- assumes basis_finite: "basis ` {dimension TYPE('a)..} = {0}"
+abbreviation dimension :: "('a::euclidean_space) itself \<Rightarrow> nat" where
+ "dimension TYPE('a) \<equiv> card (Basis :: 'a set)"
syntax "_type_dimension" :: "type => nat" ("(1DIM/(1'(_')))")
translations "DIM('t)" == "CONST dimension (TYPE('t))"
-lemma (in euclidean_space) norm_Basis: "u \<in> Basis \<Longrightarrow> norm u = 1"
+lemma (in euclidean_space) norm_Basis[simp]: "u \<in> Basis \<Longrightarrow> norm u = 1"
unfolding norm_eq_sqrt_inner by (simp add: inner_Basis)
+lemma (in euclidean_space) inner_same_Basis[simp]: "u \<in> Basis \<Longrightarrow> inner u u = 1"
+ by (simp add: inner_Basis)
+
+lemma (in euclidean_space) inner_not_same_Basis: "u \<in> Basis \<Longrightarrow> v \<in> Basis \<Longrightarrow> u \<noteq> v \<Longrightarrow> inner u v = 0"
+ by (simp add: inner_Basis)
+
lemma (in euclidean_space) sgn_Basis: "u \<in> Basis \<Longrightarrow> sgn u = u"
- unfolding sgn_div_norm by (simp add: norm_Basis scaleR_one)
+ unfolding sgn_div_norm by (simp add: scaleR_one)
lemma (in euclidean_space) Basis_zero [simp]: "0 \<notin> Basis"
proof
@@ -51,184 +51,45 @@
lemma (in euclidean_space) nonzero_Basis: "u \<in> Basis \<Longrightarrow> u \<noteq> 0"
by clarsimp
-text {* Lemmas related to @{text basis} function. *}
-
-lemma (in euclidean_space) euclidean_all_zero:
- "(\<forall>i<DIM('a). inner (basis i) x = 0) \<longleftrightarrow> (x = 0)"
- using euclidean_all_zero_iff [of x, folded image_basis]
- unfolding ball_simps by (simp add: Ball_def inner_commute)
-
-lemma (in euclidean_space) basis_zero [simp]:
- "DIM('a) \<le> i \<Longrightarrow> basis i = 0"
- using basis_finite by auto
+lemma (in euclidean_space) SOME_Basis: "(SOME i. i \<in> Basis) \<in> Basis"
+ by (metis ex_in_conv nonempty_Basis someI_ex)
-lemma (in euclidean_space) DIM_positive [intro]: "0 < DIM('a)"
- unfolding dimension_def by (simp add: card_gt_0_iff)
-
-lemma (in euclidean_space) basis_inj [simp, intro]: "inj_on basis {..<DIM('a)}"
- by (simp add: inj_on_iff_eq_card image_basis dimension_def [symmetric])
-
-lemma (in euclidean_space) basis_in_Basis [simp]:
- "basis i \<in> Basis \<longleftrightarrow> i < DIM('a)"
- by (cases "i < DIM('a)", simp add: image_basis [symmetric], simp)
-
-lemma (in euclidean_space) Basis_elim:
- assumes "u \<in> Basis" obtains i where "i < DIM('a)" and "u = basis i"
- using assms unfolding image_basis [symmetric] by fast
+lemma (in euclidean_space) inner_setsum_left_Basis[simp]:
+ "b \<in> Basis \<Longrightarrow> inner (\<Sum>i\<in>Basis. f i *\<^sub>R i) b = f b"
+ by (simp add: inner_setsum_left inner_Basis if_distrib setsum_cases)
-lemma (in euclidean_space) basis_orthonormal:
- "\<forall>i<DIM('a). \<forall>j<DIM('a).
- inner (basis i) (basis j) = (if i = j then 1 else 0)"
- apply clarify
- apply (simp add: inner_Basis)
- apply (simp add: basis_inj [unfolded inj_on_def])
- done
-
-lemma (in euclidean_space) dot_basis:
- "inner (basis i) (basis j) = (if i = j \<and> i < DIM('a) then 1 else 0)"
-proof (cases "(i < DIM('a) \<and> j < DIM('a))")
- case False
- hence "inner (basis i) (basis j) = 0" by auto
- thus ?thesis using False by auto
-next
- case True thus ?thesis using basis_orthonormal by auto
-qed
-
-lemma (in euclidean_space) basis_eq_0_iff [simp]:
- "basis i = 0 \<longleftrightarrow> DIM('a) \<le> i"
+lemma (in euclidean_space) euclidean_eqI:
+ assumes b: "\<And>b. b \<in> Basis \<Longrightarrow> inner x b = inner y b" shows "x = y"
proof -
- have "inner (basis i) (basis i) = 0 \<longleftrightarrow> DIM('a) \<le> i"
- by (simp add: dot_basis)
- thus ?thesis by simp
+ from b have "\<forall>b\<in>Basis. inner (x - y) b = 0"
+ by (simp add: inner_diff_left)
+ then show "x = y"
+ by (simp add: euclidean_all_zero_iff)
qed
-lemma (in euclidean_space) norm_basis [simp]:
- "norm (basis i) = (if i < DIM('a) then 1 else 0)"
- unfolding norm_eq_sqrt_inner dot_basis by simp
-
-lemma (in euclidean_space) basis_neq_0 [intro]:
- assumes "i<DIM('a)" shows "(basis i) \<noteq> 0"
- using assms by simp
-
-subsubsection {* Projecting components *}
-
-definition (in euclidean_space) euclidean_component (infixl "$$" 90)
- where "x $$ i = inner (basis i) x"
-
-lemma bounded_linear_euclidean_component:
- "bounded_linear (\<lambda>x. euclidean_component x i)"
- unfolding euclidean_component_def
- by (rule bounded_linear_inner_right)
-
-lemmas tendsto_euclidean_component [tendsto_intros] =
- bounded_linear.tendsto [OF bounded_linear_euclidean_component]
-
-lemmas isCont_euclidean_component [simp] =
- bounded_linear.isCont [OF bounded_linear_euclidean_component]
-
-lemma euclidean_component_zero [simp]: "0 $$ i = 0"
- unfolding euclidean_component_def by (rule inner_zero_right)
-
-lemma euclidean_component_add [simp]: "(x + y) $$ i = x $$ i + y $$ i"
- unfolding euclidean_component_def by (rule inner_add_right)
-
-lemma euclidean_component_diff [simp]: "(x - y) $$ i = x $$ i - y $$ i"
- unfolding euclidean_component_def by (rule inner_diff_right)
-
-lemma euclidean_component_minus [simp]: "(- x) $$ i = - (x $$ i)"
- unfolding euclidean_component_def by (rule inner_minus_right)
-
-lemma euclidean_component_scaleR [simp]: "(scaleR a x) $$ i = a * (x $$ i)"
- unfolding euclidean_component_def by (rule inner_scaleR_right)
-
-lemma euclidean_component_setsum [simp]: "(\<Sum>x\<in>A. f x) $$ i = (\<Sum>x\<in>A. f x $$ i)"
- unfolding euclidean_component_def by (rule inner_setsum_right)
-
-lemma euclidean_eqI:
- fixes x y :: "'a::euclidean_space"
- assumes "\<And>i. i < DIM('a) \<Longrightarrow> x $$ i = y $$ i" shows "x = y"
-proof -
- from assms have "\<forall>i<DIM('a). (x - y) $$ i = 0"
- by simp
- then show "x = y"
- unfolding euclidean_component_def euclidean_all_zero by simp
-qed
-
-lemma euclidean_eq:
- fixes x y :: "'a::euclidean_space"
- shows "x = y \<longleftrightarrow> (\<forall>i<DIM('a). x $$ i = y $$ i)"
+lemma (in euclidean_space) euclidean_eq_iff:
+ "x = y \<longleftrightarrow> (\<forall>b\<in>Basis. inner x b = inner y b)"
by (auto intro: euclidean_eqI)
-lemma (in euclidean_space) basis_component [simp]:
- "basis i $$ j = (if i = j \<and> i < DIM('a) then 1 else 0)"
- unfolding euclidean_component_def dot_basis by auto
-
-lemma (in euclidean_space) basis_at_neq_0 [intro]:
- "i < DIM('a) \<Longrightarrow> basis i $$ i \<noteq> 0"
- by simp
-
-lemma (in euclidean_space) euclidean_component_ge [simp]:
- assumes "i \<ge> DIM('a)" shows "x $$ i = 0"
- unfolding euclidean_component_def basis_zero[OF assms] by simp
+lemma (in euclidean_space) euclidean_representation_setsum:
+ "(\<Sum>i\<in>Basis. f i *\<^sub>R i) = b \<longleftrightarrow> (\<forall>i\<in>Basis. f i = inner b i)"
+ by (subst euclidean_eq_iff) simp
-lemmas euclidean_simps =
- euclidean_component_add
- euclidean_component_diff
- euclidean_component_scaleR
- euclidean_component_minus
- euclidean_component_setsum
- basis_component
-
-lemma euclidean_representation:
- fixes x :: "'a::euclidean_space"
- shows "x = (\<Sum>i<DIM('a). (x$$i) *\<^sub>R basis i)"
- apply (rule euclidean_eqI)
- apply (simp add: if_distrib setsum_delta cong: if_cong)
- done
-
-subsubsection {* Binder notation for vectors *}
-
-definition (in euclidean_space) Chi (binder "\<chi>\<chi> " 10) where
- "(\<chi>\<chi> i. f i) = (\<Sum>i<DIM('a). f i *\<^sub>R basis i)"
+lemma (in euclidean_space) euclidean_representation: "(\<Sum>b\<in>Basis. inner x b *\<^sub>R b) = x"
+ unfolding euclidean_representation_setsum by simp
-lemma euclidean_lambda_beta [simp]:
- "((\<chi>\<chi> i. f i)::'a::euclidean_space) $$ j = (if j < DIM('a) then f j else 0)"
- by (auto simp: Chi_def if_distrib setsum_cases intro!: setsum_cong)
-
-lemma euclidean_lambda_beta':
- "j < DIM('a) \<Longrightarrow> ((\<chi>\<chi> i. f i)::'a::euclidean_space) $$ j = f j"
- by simp
-
-lemma euclidean_lambda_beta'':"(\<forall>j < DIM('a::euclidean_space). P j (((\<chi>\<chi> i. f i)::'a) $$ j)) \<longleftrightarrow>
- (\<forall>j < DIM('a::euclidean_space). P j (f j))" by auto
-
-lemma euclidean_beta_reduce[simp]:
- "(\<chi>\<chi> i. x $$ i) = (x::'a::euclidean_space)"
- by (simp add: euclidean_eq)
-
-lemma euclidean_lambda_beta_0[simp]:
- "((\<chi>\<chi> i. f i)::'a::euclidean_space) $$ 0 = f 0"
- by (simp add: DIM_positive)
+lemma (in euclidean_space) choice_Basis_iff:
+ fixes P :: "'a \<Rightarrow> real \<Rightarrow> bool"
+ shows "(\<forall>i\<in>Basis. \<exists>x. P i x) \<longleftrightarrow> (\<exists>x. \<forall>i\<in>Basis. P i (inner x i))"
+ unfolding bchoice_iff
+proof safe
+ fix f assume "\<forall>i\<in>Basis. P i (f i)"
+ then show "\<exists>x. \<forall>i\<in>Basis. P i (inner x i)"
+ by (auto intro!: exI[of _ "\<Sum>i\<in>Basis. f i *\<^sub>R i"])
+qed auto
-lemma euclidean_inner:
- "inner x (y::'a) = (\<Sum>i<DIM('a::euclidean_space). (x $$ i) * (y $$ i))"
- by (subst (1 2) euclidean_representation,
- simp add: inner_setsum_left inner_setsum_right
- dot_basis if_distrib setsum_cases mult_commute)
-
-lemma euclidean_dist_l2:
- fixes x y :: "'a::euclidean_space"
- shows "dist x y = setL2 (\<lambda>i. dist (x $$ i) (y $$ i)) {..<DIM('a)}"
- unfolding dist_norm norm_eq_sqrt_inner setL2_def
- by (simp add: euclidean_inner power2_eq_square)
-
-lemma component_le_norm: "\<bar>x$$i\<bar> \<le> norm (x::'a::euclidean_space)"
- unfolding euclidean_component_def
- by (rule order_trans [OF Cauchy_Schwarz_ineq2]) simp
-
-lemma dist_nth_le: "dist (x $$ i) (y $$ i) \<le> dist x (y::'a::euclidean_space)"
- unfolding euclidean_dist_l2 [where 'a='a]
- by (cases "i < DIM('a)", rule member_le_setL2, auto)
+lemma DIM_positive: "0 < DIM('a::euclidean_space)"
+ by (simp add: card_gt_0_iff)
subsection {* Subclass relationships *}
@@ -239,11 +100,13 @@
assume "open {x}"
then obtain e where "0 < e" and e: "\<forall>y. dist y x < e \<longrightarrow> y = x"
unfolding open_dist by fast
- def y \<equiv> "x + scaleR (e/2) (sgn (basis 0))"
+ def y \<equiv> "x + scaleR (e/2) (SOME b. b \<in> Basis)"
+ have [simp]: "(SOME b. b \<in> Basis) \<in> Basis"
+ by (rule someI_ex) (auto simp: ex_in_conv)
from `0 < e` have "y \<noteq> x"
- unfolding y_def by (simp add: sgn_zero_iff DIM_positive)
+ unfolding y_def by (auto intro!: nonzero_Basis)
from `0 < e` have "dist y x < e"
- unfolding y_def by (simp add: dist_norm norm_sgn)
+ unfolding y_def by (simp add: dist_norm norm_Basis)
from `y \<noteq> x` and `dist y x < e` show "False"
using e by simp
qed
@@ -256,23 +119,17 @@
instantiation real :: euclidean_space
begin
-definition
- "Basis = {1::real}"
-
-definition
- "dimension (t::real itself) = 1"
-
-definition [simp]:
- "basis i = (if i = 0 then 1 else (0::real))"
-
-lemma DIM_real [simp]: "DIM(real) = 1"
- by (rule dimension_real_def)
+definition
+ [simp]: "Basis = {1::real}"
instance
by default (auto simp add: Basis_real_def)
end
+lemma DIM_real[simp]: "DIM(real) = 1"
+ by simp
+
subsubsection {* Type @{typ complex} *}
instantiation complex :: euclidean_space
@@ -281,20 +138,13 @@
definition Basis_complex_def:
"Basis = {1, ii}"
-definition
- "dimension (t::complex itself) = 2"
-
-definition
- "basis i = (if i = 0 then 1 else if i = 1 then ii else 0)"
-
instance
- by default (auto simp add: Basis_complex_def dimension_complex_def
- basis_complex_def intro: complex_eqI split: split_if_asm)
+ by default (auto simp add: Basis_complex_def intro: complex_eqI split: split_if_asm)
end
lemma DIM_complex[simp]: "DIM(complex) = 2"
- by (rule dimension_complex_def)
+ unfolding Basis_complex_def by simp
subsubsection {* Type @{typ "'a \<times> 'b"} *}
@@ -304,12 +154,6 @@
definition
"Basis = (\<lambda>u. (u, 0)) ` Basis \<union> (\<lambda>v. (0, v)) ` Basis"
-definition
- "dimension (t::('a \<times> 'b) itself) = DIM('a) + DIM('b)"
-
-definition
- "basis i = (if i < DIM('a) then (basis i, 0) else (0, basis (i - DIM('a))))"
-
instance proof
show "(Basis :: ('a \<times> 'b) set) \<noteq> {}"
unfolding Basis_prod_def by simp
@@ -327,20 +171,12 @@
show "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> x = 0"
unfolding Basis_prod_def ball_Un ball_simps
by (simp add: inner_prod_def prod_eq_iff euclidean_all_zero_iff)
-next
- show "DIM('a \<times> 'b) = card (Basis :: ('a \<times> 'b) set)"
- unfolding dimension_prod_def Basis_prod_def
- by (simp add: card_Un_disjoint disjoint_iff_not_equal
- card_image inj_on_def dimension_def)
-next
- show "basis ` {..<DIM('a \<times> 'b)} = (Basis :: ('a \<times> 'b) set)"
- by (auto simp add: Basis_prod_def dimension_prod_def basis_prod_def
- image_def elim!: Basis_elim)
-next
- show "basis ` {DIM('a \<times> 'b)..} = {0::('a \<times> 'b)}"
- by (auto simp add: dimension_prod_def basis_prod_def prod_eq_iff image_def)
qed
+lemma DIM_prod[simp]: "DIM('a \<times> 'b) = DIM('a) + DIM('b)"
+ unfolding Basis_prod_def
+ by (subst card_Un_disjoint) (auto intro!: card_image arg_cong2[where f="op +"] inj_onI)
+
end
end
--- a/src/HOL/Multivariate_Analysis/Fashoda.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Fashoda.thy Fri Dec 14 15:46:01 2012 +0100
@@ -7,6 +7,14 @@
imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space
begin
+(* move *)
+
+lemma cart_eq_inner_axis: "a $ i = a \<bullet> axis i 1"
+ by (simp add: inner_axis)
+
+lemma axis_in_Basis: "a \<in> Basis \<Longrightarrow> axis i a \<in> Basis"
+ by (auto simp add: Basis_vec_def axis_eq_axis)
+
subsection {*Fashoda meet theorem. *}
lemma infnorm_2: "infnorm (x::real^2) = max (abs(x$1)) (abs(x$2))"
@@ -30,7 +38,7 @@
have lem1:"\<forall>z::real^2. infnorm(negatex z) = infnorm z"
unfolding negatex_def infnorm_2 vector_2 by auto
have lem2:"\<forall>z. z\<noteq>0 \<longrightarrow> infnorm(sqprojection z) = 1" unfolding sqprojection_def
- unfolding infnorm_mul[unfolded smult_conv_scaleR] unfolding abs_inverse real_abs_infnorm
+ unfolding infnorm_mul[unfolded scalar_mult_eq_scaleR] unfolding abs_inverse real_abs_infnorm
apply(subst infnorm_eq_0[THEN sym]) by auto
let ?F = "(\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w)"
have *:"\<And>i. (\<lambda>x::real^2. x $ i) ` {- 1..1} = {- 1..1::real}"
@@ -133,12 +141,6 @@
apply(rule_tac x="iscale s" in bexI) prefer 3 apply(rule_tac x="iscale t" in bexI)
using isc[unfolded subset_eq, rule_format] by auto qed
-(* move *)
-lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\<forall>i. a$i < b$i \<and> u$i < v$i"
- shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x"
- unfolding interval_bij_cart split_conv vec_eq_iff vec_lambda_beta
- apply(rule,insert assms,erule_tac x=i in allE) by auto
-
lemma fashoda: fixes b::"real^2"
assumes "path f" "path g" "path_image f \<subseteq> {a..b}" "path_image g \<subseteq> {a..b}"
"(pathstart f)$1 = a$1" "(pathfinish f)$1 = b$1"
@@ -184,8 +186,10 @@
"(interval_bij (a, b) (- 1, 1) \<circ> f) 1 $ 1 = 1"
"(interval_bij (a, b) (- 1, 1) \<circ> g) 0 $ 2 = -1"
"(interval_bij (a, b) (- 1, 1) \<circ> g) 1 $ 2 = 1"
- unfolding interval_bij_cart vector_component_simps o_def split_conv
- unfolding assms[unfolded pathstart_def pathfinish_def] using as by auto qed note z=this
+ 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 note z=this
from z(1) guess zf unfolding image_iff .. note zf=this
from z(2) guess zg unfolding image_iff .. note zg=this
have *:"\<forall>i. (- 1) $ i < (1::real^2) $ i \<and> a $ i < b $ i" unfolding forall_2 using as by auto
@@ -201,7 +205,7 @@
proof-
let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"
{ presume "?L \<Longrightarrow> ?R" "?R \<Longrightarrow> ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq
- unfolding vec_eq_iff forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast }
+ unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[THEN sym] vector_component_simps by blast }
{ assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this
{ fix b a assume "b + u * a > a + u * b"
hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps)
@@ -225,7 +229,7 @@
proof-
let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"
{ presume "?L \<Longrightarrow> ?R" "?R \<Longrightarrow> ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq
- unfolding vec_eq_iff forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast }
+ unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[THEN sym] vector_component_simps by blast }
{ assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this
{ fix b a assume "b + u * a > a + u * b"
hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps)
--- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Fri Dec 14 15:46:01 2012 +0100
@@ -504,96 +504,11 @@
apply (simp add: axis_def)
done
-text {* A bijection between @{text "'n::finite"} and @{text "{..<CARD('n)}"} *}
-
-definition vec_bij_nat :: "nat \<Rightarrow> ('n::finite)" where
- "vec_bij_nat = (SOME p. bij_betw p {..<CARD('n)} (UNIV::'n set) )"
-
-abbreviation "\<pi> \<equiv> vec_bij_nat"
-definition "\<pi>' = inv_into {..<CARD('n)} (\<pi>::nat \<Rightarrow> ('n::finite))"
-
-lemma bij_betw_pi:
- "bij_betw \<pi> {..<CARD('n::finite)} (UNIV::('n::finite) set)"
- using ex_bij_betw_nat_finite[of "UNIV::'n set"]
- by (auto simp: vec_bij_nat_def atLeast0LessThan
- intro!: someI_ex[of "\<lambda>x. bij_betw x {..<CARD('n)} (UNIV::'n set)"])
-
-lemma bij_betw_pi'[intro]: "bij_betw \<pi>' (UNIV::'n set) {..<CARD('n::finite)}"
- using bij_betw_inv_into[OF bij_betw_pi] unfolding \<pi>'_def by auto
-
-lemma pi'_inj[intro]: "inj \<pi>'"
- using bij_betw_pi' unfolding bij_betw_def by auto
-
-lemma pi'_range[intro]: "\<And>i::'n. \<pi>' i < CARD('n::finite)"
- using bij_betw_pi' unfolding bij_betw_def by auto
-
-lemma pi_pi'[simp]: "\<And>i::'n::finite. \<pi> (\<pi>' i) = i"
- using bij_betw_pi by (auto intro!: f_inv_into_f simp: \<pi>'_def bij_betw_def)
-
-lemma pi'_pi[simp]: "\<And>i. i\<in>{..<CARD('n::finite)} \<Longrightarrow> \<pi>' (\<pi> i::'n) = i"
- using bij_betw_pi by (auto intro!: inv_into_f_eq simp: \<pi>'_def bij_betw_def)
-
-lemma pi_pi'_alt[simp]: "\<And>i. i<CARD('n::finite) \<Longrightarrow> \<pi>' (\<pi> i::'n) = i"
- by auto
-
-lemma pi_inj_on: "inj_on (\<pi>::nat\<Rightarrow>'n::finite) {..<CARD('n)}"
- using bij_betw_pi[where 'n='n] by (simp add: bij_betw_def)
-
instantiation vec :: (euclidean_space, finite) euclidean_space
begin
definition "Basis = (\<Union>i. \<Union>u\<in>Basis. {axis i u})"
-definition "dimension (t :: ('a ^ 'b) itself) = CARD('b) * DIM('a)"
-
-definition "basis i =
- (if i < (CARD('b) * DIM('a))
- then axis (\<pi>(i div DIM('a))) (basis (i mod DIM('a)))
- else 0)"
-
-lemma basis_eq:
- assumes "i < CARD('b)" and "j < DIM('a)"
- shows "basis (j + i * DIM('a)) = axis (\<pi> i) (basis j)"
-proof -
- have "j + i * DIM('a) < DIM('a) * (i + 1)" using assms by (auto simp: field_simps)
- also have "\<dots> \<le> DIM('a) * CARD('b)" using assms unfolding mult_le_cancel1 by auto
- finally show ?thesis
- unfolding basis_vec_def using assms by (auto simp: vec_eq_iff not_less field_simps)
-qed
-
-lemma basis_eq_pi':
- assumes "j < DIM('a)"
- shows "basis (j + \<pi>' i * DIM('a)) $ k = (if k = i then basis j else 0)"
- apply (subst basis_eq)
- using pi'_range assms by (simp_all add: axis_def)
-
-lemma split_times_into_modulo[consumes 1]:
- fixes k :: nat
- assumes "k < A * B"
- obtains i j where "i < A" and "j < B" and "k = j + i * B"
-proof
- have "A * B \<noteq> 0"
- proof assume "A * B = 0" with assms show False by simp qed
- hence "0 < B" by auto
- thus "k mod B < B" using `0 < B` by auto
-next
- have "k div B * B \<le> k div B * B + k mod B" by (rule le_add1)
- also have "... < A * B" using assms by simp
- finally show "k div B < A" by auto
-qed simp
-
-lemma linear_less_than_times:
- fixes i j A B :: nat assumes "i < B" "j < A"
- shows "j + i * A < B * A"
-proof -
- have "i * A + j < (Suc i)*A" using `j < A` by simp
- also have "\<dots> \<le> B * A" using `i < B` unfolding mult_le_cancel2 by simp
- finally show ?thesis by simp
-qed
-
-lemma DIM_cart[simp]: "DIM('a^'b) = CARD('b) * DIM('a)"
- by (rule dimension_vec_def)
-
instance proof
show "(Basis :: ('a ^ 'b) set) \<noteq> {}"
unfolding Basis_vec_def by simp
@@ -611,27 +526,17 @@
show "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> x = 0"
unfolding Basis_vec_def
by (simp add: inner_axis euclidean_all_zero_iff vec_eq_iff)
-next
- show "DIM('a ^ 'b) = card (Basis :: ('a ^ 'b) set)"
- unfolding Basis_vec_def dimension_vec_def dimension_def
- by (simp add: card_UN_disjoint [unfolded disjoint_iff_not_equal]
- axis_eq_axis nonzero_Basis)
-next
- show "basis ` {..<DIM('a ^ 'b)} = (Basis :: ('a ^ 'b) set)"
- unfolding Basis_vec_def
- apply auto
- apply (erule split_times_into_modulo)
- apply (simp add: basis_eq axis_eq_axis)
- apply (erule Basis_elim)
- apply (simp add: image_def basis_vec_def axis_eq_axis)
- apply (rule rev_bexI, simp)
- apply (erule linear_less_than_times [OF pi'_range])
+qed
+
+lemma DIM_cart[simp]: "DIM('a^'b) = CARD('b) * DIM('a)"
+ apply (simp add: Basis_vec_def)
+ apply (subst card_UN_disjoint)
+ apply simp
apply simp
- done
-next
- show "basis ` {DIM('a ^ 'b)..} = {0::'a ^ 'b}"
- by (auto simp add: image_def basis_vec_def)
-qed
+ apply (auto simp: axis_eq_axis) [1]
+ apply (subst card_UN_disjoint)
+ apply (auto simp: axis_eq_axis)
+ done
end
--- a/src/HOL/Multivariate_Analysis/Integration.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Fri Dec 14 15:46:01 2012 +0100
@@ -71,12 +71,8 @@
apply (auto simp: isUb_def setle_def)
done
-lemma bounded_linear_component [intro]: "bounded_linear (\<lambda>x::'a::euclidean_space. x $$ k)"
- apply (rule bounded_linearI[where K=1])
- using component_le_norm[of _ k]
- unfolding real_norm_def
- apply auto
- done
+lemma bounded_linear_component [intro]: "bounded_linear (\<lambda>x::'a::euclidean_space. x \<bullet> k)"
+ by (rule bounded_linear_inner_left)
lemma transitive_stepwise_lt_eq:
assumes "(\<And>x y z::nat. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z)"
@@ -162,16 +158,10 @@
subsection {* Some useful lemmas about intervals. *}
-abbreviation One where "One \<equiv> ((\<chi>\<chi> i. 1)::_::ordered_euclidean_space)"
-
-lemma empty_as_interval: "{} = {One..0}"
- apply (rule set_eqI, rule)
- defer
- unfolding mem_interval
- using UNIV_witness[where 'a='n]
- apply (erule_tac exE, rule_tac x = x in allE)
- apply auto
- done
+abbreviation One where "One \<equiv> ((\<Sum>Basis)::_::euclidean_space)"
+
+lemma empty_as_interval: "{} = {One..(0::'a::ordered_euclidean_space)}"
+ by (auto simp: set_eq_iff eucl_le[where 'a='a] intro!: bexI[OF _ SOME_Basis])
lemma interior_subset_union_intervals:
assumes "i = {a..b::'a::ordered_euclidean_space}" "j = {c..d}"
@@ -255,17 +245,17 @@
done
next
case False
- then obtain k where "x$$k \<le> a$$k \<or> x$$k \<ge> b$$k" and k:"k<DIM('a)"
+ then obtain k where "x\<bullet>k \<le> a\<bullet>k \<or> x\<bullet>k \<ge> b\<bullet>k" and k:"k\<in>Basis"
unfolding mem_interval by (auto simp add: not_less)
- hence "x$$k = a$$k \<or> x$$k = b$$k"
+ hence "x\<bullet>k = a\<bullet>k \<or> x\<bullet>k = b\<bullet>k"
using True unfolding ab and mem_interval
- apply (erule_tac x = k in allE)
+ apply (erule_tac x = k in ballE)
apply auto
done
hence "\<exists>x. ball x (e/2) \<subseteq> s \<inter> (\<Union>f)"
proof (erule_tac disjE)
- let ?z = "x - (e/2) *\<^sub>R basis k"
- assume as: "x$$k = a$$k"
+ let ?z = "x - (e/2) *\<^sub>R k"
+ assume as: "x\<bullet>k = a\<bullet>k"
have "ball ?z (e / 2) \<inter> i = {}"
apply (rule ccontr)
unfolding ex_in_conv[THEN sym]
@@ -273,15 +263,12 @@
fix y
assume "y \<in> ball ?z (e / 2) \<inter> i"
hence "dist ?z y < e/2" and yi:"y\<in>i" by auto
- hence "\<bar>(?z - y) $$ k\<bar> < e/2"
- using component_le_norm[of "?z - y" k] unfolding dist_norm by auto
- hence "y$$k < a$$k"
- using e[THEN conjunct1] k by (auto simp add: field_simps as)
+ hence "\<bar>(?z - y) \<bullet> k\<bar> < e/2"
+ using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto
+ hence "y\<bullet>k < a\<bullet>k"
+ using e[THEN conjunct1] k by (auto simp add: field_simps as inner_Basis inner_simps)
hence "y \<notin> i"
- unfolding ab mem_interval not_all
- apply (rule_tac x=k in exI)
- using k apply auto
- done
+ unfolding ab mem_interval by (auto intro!: bexI[OF _ k])
thus False using yi by auto
qed
moreover
@@ -290,10 +277,10 @@
proof
fix y
assume as: "y\<in> ball ?z (e/2)"
- have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y - (e / 2) *\<^sub>R basis k)"
+ have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y - (e / 2) *\<^sub>R k)"
apply -
- apply (rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R basis k"])
- unfolding norm_scaleR norm_basis
+ apply (rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R k"])
+ unfolding norm_scaleR norm_Basis[OF k]
apply auto
done
also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2"
@@ -310,8 +297,8 @@
apply auto
done
next
- let ?z = "x + (e/2) *\<^sub>R basis k"
- assume as: "x$$k = b$$k"
+ let ?z = "x + (e/2) *\<^sub>R k"
+ assume as: "x\<bullet>k = b\<bullet>k"
have "ball ?z (e / 2) \<inter> i = {}"
apply (rule ccontr)
unfolding ex_in_conv[THEN sym]
@@ -319,15 +306,12 @@
fix y
assume "y \<in> ball ?z (e / 2) \<inter> i"
hence "dist ?z y < e/2" and yi:"y\<in>i" by auto
- hence "\<bar>(?z - y) $$ k\<bar> < e/2"
- using component_le_norm[of "?z - y" k] unfolding dist_norm by auto
- hence "y$$k > b$$k"
- using e[THEN conjunct1] k by(auto simp add:field_simps as)
+ hence "\<bar>(?z - y) \<bullet> k\<bar> < e/2"
+ using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto
+ hence "y\<bullet>k > b\<bullet>k"
+ using e[THEN conjunct1] k by(auto simp add:field_simps inner_simps inner_Basis as)
hence "y \<notin> i"
- unfolding ab mem_interval not_all
- using k apply (rule_tac x=k in exI)
- apply auto
- done
+ unfolding ab mem_interval by (auto intro!: bexI[OF _ k])
thus False using yi by auto
qed
moreover
@@ -336,11 +320,11 @@
proof
fix y
assume as: "y\<in> ball ?z (e/2)"
- have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y + (e / 2) *\<^sub>R basis k)"
+ have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y + (e / 2) *\<^sub>R k)"
apply -
- apply(rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R basis k"])
+ apply(rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R k"])
unfolding norm_scaleR
- apply auto
+ apply (auto simp: k)
done
also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2"
apply (rule add_strict_left_mono)
@@ -383,61 +367,25 @@
subsection {* Bounds on intervals where they exist. *}
-definition "interval_upperbound (s::('a::ordered_euclidean_space) set) =
- ((\<chi>\<chi> i. Sup {a. \<exists>x\<in>s. x$$i = a})::'a)"
-
-definition "interval_lowerbound (s::('a::ordered_euclidean_space) set) =
- ((\<chi>\<chi> i. Inf {a. \<exists>x\<in>s. x$$i = a})::'a)"
+definition interval_upperbound :: "('a::ordered_euclidean_space) set \<Rightarrow> 'a" where
+ "interval_upperbound s = (\<Sum>i\<in>Basis. Sup {a. \<exists>x\<in>s. x\<bullet>i = a} *\<^sub>R i)"
+
+definition interval_lowerbound :: "('a::ordered_euclidean_space) set \<Rightarrow> 'a" where
+ "interval_lowerbound s = (\<Sum>i\<in>Basis. Inf {a. \<exists>x\<in>s. x\<bullet>i = a} *\<^sub>R i)"
lemma interval_upperbound[simp]:
- assumes "\<forall>i<DIM('a::ordered_euclidean_space). a$$i \<le> (b::'a)$$i"
- shows "interval_upperbound {a..b} = b"
- using assms
- unfolding interval_upperbound_def
- apply (subst euclidean_eq[where 'a='a])
- apply safe
- unfolding euclidean_lambda_beta'
- apply (erule_tac x=i in allE)
- apply (rule Sup_unique)
- unfolding setle_def
- apply rule
- unfolding mem_Collect_eq
- apply (erule bexE)
- unfolding mem_interval
- defer
- apply (rule, rule)
- apply (rule_tac x="b$$i" in bexI)
- defer
- unfolding mem_Collect_eq
- apply (rule_tac x=b in bexI)
- unfolding mem_interval
- using assms apply auto
- done
+ "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
+ interval_upperbound {a..b} = (b::'a::ordered_euclidean_space)"
+ unfolding interval_upperbound_def euclidean_representation_setsum
+ by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setle_def
+ intro!: Sup_unique)
lemma interval_lowerbound[simp]:
- assumes "\<forall>i<DIM('a::ordered_euclidean_space). a$$i \<le> (b::'a)$$i"
- shows "interval_lowerbound {a..b} = a"
- using assms
- unfolding interval_lowerbound_def
- apply (subst euclidean_eq[where 'a='a])
- apply safe
- unfolding euclidean_lambda_beta'
- apply (erule_tac x=i in allE)
- apply (rule Inf_unique)
- unfolding setge_def
- apply rule
- unfolding mem_Collect_eq
- apply (erule bexE)
- unfolding mem_interval
- defer
- apply (rule, rule)
- apply (rule_tac x = "a$$i" in bexI)
- defer
- unfolding mem_Collect_eq
- apply (rule_tac x=a in bexI)
- unfolding mem_interval
- using assms apply auto
- done
+ "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
+ interval_lowerbound {a..b} = (a::'a::ordered_euclidean_space)"
+ unfolding interval_lowerbound_def euclidean_representation_setsum
+ by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setge_def
+ intro!: Inf_unique)
lemmas interval_bounds = interval_upperbound interval_lowerbound
@@ -449,15 +397,15 @@
subsection {* Content (length, area, volume...) of an interval. *}
definition "content (s::('a::ordered_euclidean_space) set) =
- (if s = {} then 0 else (\<Prod>i<DIM('a). (interval_upperbound s)$$i - (interval_lowerbound s)$$i))"
-
-lemma interval_not_empty:"\<forall>i<DIM('a). a$$i \<le> b$$i \<Longrightarrow> {a..b::'a::ordered_euclidean_space} \<noteq> {}"
+ (if s = {} then 0 else (\<Prod>i\<in>Basis. (interval_upperbound s)\<bullet>i - (interval_lowerbound s)\<bullet>i))"
+
+lemma interval_not_empty:"\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> {a..b::'a::ordered_euclidean_space} \<noteq> {}"
unfolding interval_eq_empty unfolding not_ex not_less by auto
lemma content_closed_interval:
fixes a::"'a::ordered_euclidean_space"
- assumes "\<forall>i<DIM('a). a$$i \<le> b$$i"
- shows "content {a..b} = (\<Prod>i<DIM('a). b$$i - a$$i)"
+ assumes "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
+ shows "content {a..b} = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
using interval_not_empty[OF assms]
unfolding content_def interval_upperbound[OF assms] interval_lowerbound[OF assms]
by auto
@@ -465,7 +413,7 @@
lemma content_closed_interval':
fixes a::"'a::ordered_euclidean_space"
assumes "{a..b}\<noteq>{}"
- shows "content {a..b} = (\<Prod>i<DIM('a). b$$i - a$$i)"
+ shows "content {a..b} = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
apply (rule content_closed_interval)
using assms unfolding interval_ne_empty
apply assumption
@@ -482,13 +430,13 @@
lemma content_singleton[simp]: "content {a} = 0"
proof -
have "content {a .. a} = 0"
- by (subst content_closed_interval) auto
+ 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"
proof -
- have *: "\<forall>i<DIM('a). (0::'a)$$i \<le> (One::'a)$$i" by auto
+ have *: "\<forall>i\<in>Basis. (0::'a)\<bullet>i \<le> (One::'a)\<bullet>i" by auto
have "0 \<in> {0..One::'a}" unfolding mem_interval by auto
thus ?thesis unfolding content_def interval_bounds[OF *] using setprod_1 by auto
qed
@@ -498,12 +446,12 @@
shows "0 \<le> content {a..b}"
proof (cases "{a..b} = {}")
case False
- hence *: "\<forall>i<DIM('a). a $$ i \<le> b $$ i" unfolding interval_ne_empty .
- have "(\<Prod>i<DIM('a). interval_upperbound {a..b} $$ i - interval_lowerbound {a..b} $$ i) \<ge> 0"
+ hence *: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i" unfolding interval_ne_empty .
+ have "(\<Prod>i\<in>Basis. interval_upperbound {a..b} \<bullet> i - interval_lowerbound {a..b} \<bullet> i) \<ge> 0"
apply (rule setprod_nonneg)
unfolding interval_bounds[OF *]
using *
- apply (erule_tac x=x in allE)
+ apply (erule_tac x=x in ballE)
apply auto
done
thus ?thesis unfolding content_def by (auto simp del:interval_bounds')
@@ -511,75 +459,59 @@
lemma content_pos_lt:
fixes a::"'a::ordered_euclidean_space"
- assumes "\<forall>i<DIM('a). a$$i < b$$i"
+ assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"
shows "0 < content {a..b}"
proof -
- have help_lemma1: "\<forall>i<DIM('a). a$$i < b$$i \<Longrightarrow> \<forall>i<DIM('a). a$$i \<le> ((b$$i)::real)"
- apply (rule, erule_tac x=i in allE)
+ have help_lemma1: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i \<Longrightarrow> \<forall>i\<in>Basis. a\<bullet>i \<le> ((b\<bullet>i)::real)"
+ apply (rule, erule_tac x=i in ballE)
apply auto
done
show ?thesis unfolding content_closed_interval[OF help_lemma1[OF assms]]
apply(rule setprod_pos)
- using assms apply (erule_tac x=x in allE)
+ using assms apply (erule_tac x=x in ballE)
apply auto
done
qed
-lemma content_eq_0: "content{a..b::'a::ordered_euclidean_space} = 0 \<longleftrightarrow> (\<exists>i<DIM('a). b$$i \<le> a$$i)"
+lemma content_eq_0: "content{a..b::'a::ordered_euclidean_space} = 0 \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i)"
proof (cases "{a..b} = {}")
case True
thus ?thesis
unfolding content_def if_P[OF True]
unfolding interval_eq_empty
apply -
- apply (rule, erule exE)
- apply (rule_tac x = i in exI)
+ apply (rule, erule bexE)
+ apply (rule_tac x = i in bexI)
apply auto
done
next
case False
from this[unfolded interval_eq_empty not_ex not_less]
- have as: "\<forall>i<DIM('a). b $$ i \<ge> a $$ i" by fastforce
+ have as: "\<forall>i\<in>Basis. b \<bullet> i \<ge> a \<bullet> i" by fastforce
show ?thesis
- unfolding content_def if_not_P[OF False] setprod_zero_iff[OF finite_lessThan]
- apply rule
- apply (erule_tac[!] exE bexE)
- unfolding interval_bounds[OF as]
- apply (rule_tac x=x in exI)
- defer
- apply (rule_tac x=i in bexI)
- using as apply (erule_tac x=i in allE)
- apply auto
- done
+ unfolding content_def if_not_P[OF False] setprod_zero_iff[OF finite_Basis]
+ using as
+ by (auto intro!: bexI)
qed
lemma cond_cases:"(P \<Longrightarrow> Q x) \<Longrightarrow> (\<not> P \<Longrightarrow> Q y) \<Longrightarrow> Q (if P then x else y)" by auto
lemma content_closed_interval_cases:
"content {a..b::'a::ordered_euclidean_space} =
- (if \<forall>i<DIM('a). a$$i \<le> b$$i then setprod (\<lambda>i. b$$i - a$$i) {..<DIM('a)} else 0)"
- apply (rule cond_cases)
- apply (rule content_closed_interval)
- unfolding content_eq_0 not_all not_le
- defer
- apply (erule exE,rule_tac x=x in exI)
- apply auto
- done
+ (if \<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i then setprod (\<lambda>i. b\<bullet>i - a\<bullet>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 \<longleftrightarrow> interior({a..b}) = {}"
unfolding content_eq_0 interior_closed_interval interval_eq_empty by auto
-(*lemma content_eq_0_1: "content {a..b::real^1} = 0 \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
- unfolding content_eq_0 by auto*)
-
-lemma content_pos_lt_eq: "0 < content {a..b::'a::ordered_euclidean_space} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i < b$$i)"
+lemma content_pos_lt_eq: "0 < content {a..b::'a::ordered_euclidean_space} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)"
apply rule
defer
apply (rule content_pos_lt, assumption)
proof -
assume "0 < content {a..b}"
hence "content {a..b} \<noteq> 0" by auto
- thus "\<forall>i<DIM('a). a$$i < b$$i"
+ thus "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"
unfolding content_eq_0 not_ex not_le by fastforce
qed
@@ -593,20 +525,20 @@
thus ?thesis using content_pos_le[of c d] by auto
next
case False
- hence ab_ne:"\<forall>i<DIM('a). a $$ i \<le> b $$ i" unfolding interval_ne_empty by auto
+ hence ab_ne:"\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i" unfolding interval_ne_empty by auto
hence ab_ab:"a\<in>{a..b}" "b\<in>{a..b}" unfolding mem_interval by auto
have "{c..d} \<noteq> {}" using assms False by auto
- hence cd_ne:"\<forall>i<DIM('a). c $$ i \<le> d $$ i" using assms unfolding interval_ne_empty by auto
+ hence cd_ne:"\<forall>i\<in>Basis. c \<bullet> i \<le> d \<bullet> i" using assms unfolding interval_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 `{c..d} \<noteq> {}`]
apply(rule setprod_mono,rule)
proof
- fix i
- assume i:"i\<in>{..<DIM('a)}"
- show "0 \<le> b $$ i - a $$ i" using ab_ne[THEN spec[where x=i]] i by auto
- show "b $$ i - a $$ i \<le> d $$ i - c $$ i"
+ fix i :: 'a
+ assume i:"i\<in>Basis"
+ show "0 \<le> b \<bullet> i - a \<bullet> i" using ab_ne[THEN bspec, OF i] i by auto
+ show "b \<bullet> i - a \<bullet> i \<le> d \<bullet> i - c \<bullet> i"
using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(2),of i]
using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(1),of i]
using i by auto
@@ -857,134 +789,84 @@
show "k \<noteq> {}" using k d1(3) d2(3) by auto show "\<exists>a b. k = {a..b}" using k d1(4) d2(4) by auto qed
lemma partial_division_extend_1:
- assumes "{c..d} \<subseteq> {a..b::'a::ordered_euclidean_space}" "{c..d} \<noteq> {}"
+ assumes incl: "{c..d} \<subseteq> {a..b::'a::ordered_euclidean_space}" and nonempty: "{c..d} \<noteq> {}"
obtains p where "p division_of {a..b}" "{c..d} \<in> p"
-proof- def n \<equiv> "DIM('a)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_def using DIM_positive[where 'a='a] by auto
- guess \<pi> using ex_bij_betw_nat_finite_1[OF finite_lessThan[of "DIM('a)"]] .. note \<pi>=this
- def \<pi>' \<equiv> "inv_into {1..n} \<pi>"
- have \<pi>':"bij_betw \<pi>' {..<DIM('a)} {1..n}" using bij_betw_inv_into[OF \<pi>] unfolding \<pi>'_def n_def by auto
- hence \<pi>'_i:"\<And>i. i<DIM('a) \<Longrightarrow> \<pi>' i \<in> {1..n}" unfolding bij_betw_def by auto
- have \<pi>_i:"\<And>i. i\<in>{1..n} \<Longrightarrow> \<pi> i <DIM('a)" using \<pi> unfolding bij_betw_def n_def by auto
- have \<pi>_\<pi>'[simp]:"\<And>i. i<DIM('a) \<Longrightarrow> \<pi> (\<pi>' i) = i" unfolding \<pi>'_def
- apply(rule f_inv_into_f) unfolding n_def using \<pi> unfolding bij_betw_def by auto
- have \<pi>'_\<pi>[simp]:"\<And>i. i\<in>{1..n} \<Longrightarrow> \<pi>' (\<pi> i) = i" unfolding \<pi>'_def apply(rule inv_into_f_eq)
- using \<pi> unfolding n_def bij_betw_def by auto
- have "{c..d} \<noteq> {}" using assms by auto
- let ?p1 = "\<lambda>l. {(\<chi>\<chi> i. if \<pi>' i < l then c$$i else a$$i)::'a .. (\<chi>\<chi> i. if \<pi>' i < l then d$$i else if \<pi>' i = l then c$$\<pi> l else b$$i)}"
- let ?p2 = "\<lambda>l. {(\<chi>\<chi> i. if \<pi>' i < l then c$$i else if \<pi>' i = l then d$$\<pi> l else a$$i)::'a .. (\<chi>\<chi> i. if \<pi>' i < l then d$$i else b$$i)}"
- let ?p = "{?p1 l |l. l \<in> {1..n+1}} \<union> {?p2 l |l. l \<in> {1..n+1}}"
- have abcd:"\<And>i. i<DIM('a) \<Longrightarrow> a $$ i \<le> c $$ i \<and> c$$i \<le> d$$i \<and> d $$ i \<le> b $$ i" using assms
- unfolding subset_interval interval_eq_empty by auto
- show ?thesis apply(rule that[of ?p]) apply(rule division_ofI)
- proof- have "\<And>i. i<DIM('a) \<Longrightarrow> \<pi>' i < Suc n"
- proof(rule ccontr,unfold not_less) fix i assume i:"i<DIM('a)" and "Suc n \<le> \<pi>' i"
- hence "\<pi>' i \<notin> {1..n}" by auto thus False using \<pi>' i unfolding bij_betw_def by auto
- qed hence "c = (\<chi>\<chi> i. if \<pi>' i < Suc n then c $$ i else a $$ i)"
- "d = (\<chi>\<chi> i. if \<pi>' i < Suc n then d $$ i else if \<pi>' i = n + 1 then c $$ \<pi> (n + 1) else b $$ i)"
- unfolding euclidean_eq[where 'a='a] using \<pi>' unfolding bij_betw_def by auto
- thus cdp:"{c..d} \<in> ?p" apply-apply(rule UnI1) unfolding mem_Collect_eq apply(rule_tac x="n + 1" in exI) by auto
- have "\<And>l. l\<in>{1..n+1} \<Longrightarrow> ?p1 l \<subseteq> {a..b}" "\<And>l. l\<in>{1..n+1} \<Longrightarrow> ?p2 l \<subseteq> {a..b}"
- unfolding subset_eq apply(rule_tac[!] ballI,rule_tac[!] ccontr)
- proof- fix l assume l:"l\<in>{1..n+1}" fix x assume "x\<notin>{a..b}"
- then guess i unfolding mem_interval not_all not_imp .. note i=conjunctD2[OF this]
- show "x \<in> ?p1 l \<Longrightarrow> False" "x \<in> ?p2 l \<Longrightarrow> False" unfolding mem_interval apply(erule_tac[!] x=i in allE)
- apply(case_tac[!] "\<pi>' i < l", case_tac[!] "\<pi>' i = l") using abcd[of i] i by auto
- qed moreover have "\<And>x. x \<in> {a..b} \<Longrightarrow> x \<in> \<Union>?p"
- proof- fix x assume x:"x\<in>{a..b}"
- { presume "x\<notin>{c..d} \<Longrightarrow> x \<in> \<Union>?p" thus "x \<in> \<Union>?p" using cdp by blast }
- let ?M = "{i. i\<in>{1..n+1} \<and> \<not> (c $$ \<pi> i \<le> x $$ \<pi> i \<and> x $$ \<pi> i \<le> d $$ \<pi> i)}"
- assume "x\<notin>{c..d}" then guess i0 unfolding mem_interval not_all not_imp ..
- hence "\<pi>' i0 \<in> ?M" using \<pi>' unfolding bij_betw_def by(auto intro!:le_SucI)
- hence M:"finite ?M" "?M \<noteq> {}" by auto
- def l \<equiv> "Min ?M" note l = Min_less_iff[OF M,unfolded l_def[symmetric]] Min_in[OF M,unfolded mem_Collect_eq l_def[symmetric]]
- Min_gr_iff[OF M,unfolded l_def[symmetric]]
- have "x\<in>?p1 l \<or> x\<in>?p2 l" using l(2)[THEN conjunct2] unfolding de_Morgan_conj not_le
- apply- apply(erule disjE) apply(rule disjI1) defer apply(rule disjI2)
- proof- assume as:"x $$ \<pi> l < c $$ \<pi> l"
- show "x \<in> ?p1 l" unfolding mem_interval apply safe unfolding euclidean_lambda_beta'
- proof- case goal1 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le using goal1 by auto
- thus ?case using as x[unfolded mem_interval,rule_format,of i]
- apply auto using l(3)[of "\<pi>' i"] using goal1 by(auto elim!:ballE[where x="\<pi>' i"])
- next case goal2 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le using goal2 by auto
- thus ?case using as x[unfolded mem_interval,rule_format,of i]
- apply auto using l(3)[of "\<pi>' i"] using goal2 by(auto elim!:ballE[where x="\<pi>' i"])
- qed
- next assume as:"x $$ \<pi> l > d $$ \<pi> l"
- show "x \<in> ?p2 l" unfolding mem_interval apply safe unfolding euclidean_lambda_beta'
- proof- fix i assume i:"i<DIM('a)"
- have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le using i by auto
- thus "(if \<pi>' i < l then c $$ i else if \<pi>' i = l then d $$ \<pi> l else a $$ i) \<le> x $$ i"
- "x $$ i \<le> (if \<pi>' i < l then d $$ i else b $$ i)"
- using as x[unfolded mem_interval,rule_format,of i]
- apply auto using l(3)[of "\<pi>' i"] i by(auto elim!:ballE[where x="\<pi>' i"])
- qed qed
- thus "x \<in> \<Union>?p" using l(2) by blast
- qed ultimately show "\<Union>?p = {a..b}" apply-apply(rule) defer apply(rule) by(assumption,blast)
-
- show "finite ?p" by auto
- fix k assume k:"k\<in>?p" then obtain l where l:"k = ?p1 l \<or> k = ?p2 l" "l \<in> {1..n + 1}" by auto
- show "k\<subseteq>{a..b}" apply(rule,unfold mem_interval,rule,rule)
- proof fix i x assume i:"i<DIM('a)" assume "x \<in> k" moreover have "\<pi>' i < l \<or> \<pi>' i = l \<or> \<pi>' i > l" by auto
- ultimately show "a$$i \<le> x$$i" "x$$i \<le> b$$i" using abcd[of i] using l using i
- by(auto elim!:allE[where x=i] simp add:eucl_le[where 'a='a]) (* FIXME: SLOW *)
- qed have "\<And>l. ?p1 l \<noteq> {}" "\<And>l. ?p2 l \<noteq> {}" unfolding interval_eq_empty not_ex apply(rule_tac[!] allI)
- proof- case goal1 thus ?case using abcd[of x] by auto
- next case goal2 thus ?case using abcd[of x] by auto
- qed thus "k \<noteq> {}" using k by auto
- show "\<exists>a b. k = {a..b}" using k by auto
- fix k' assume k':"k' \<in> ?p" "k \<noteq> k'" then obtain l' where l':"k' = ?p1 l' \<or> k' = ?p2 l'" "l' \<in> {1..n + 1}" by auto
- { fix k k' l l'
- assume k:"k\<in>?p" and l:"k = ?p1 l \<or> k = ?p2 l" "l \<in> {1..n + 1}"
- assume k':"k' \<in> ?p" "k \<noteq> k'" and l':"k' = ?p1 l' \<or> k' = ?p2 l'" "l' \<in> {1..n + 1}"
- assume "l \<le> l'" fix x
- have "x \<notin> interior k \<inter> interior k'"
- proof(rule,cases "l' = n+1") assume x:"x \<in> interior k \<inter> interior k'"
- case True hence "\<And>i. i<DIM('a) \<Longrightarrow> \<pi>' i < l'" using \<pi>'_i using l' by(auto simp add:less_Suc_eq_le)
- hence *:"\<And> P Q. (\<chi>\<chi> i. if \<pi>' i < l' then P i else Q i) = ((\<chi>\<chi> i. P i)::'a)" apply-apply(subst euclidean_eq) by auto
- hence k':"k' = {c..d}" using l'(1) unfolding * by auto
- have ln:"l < n + 1"
- proof(rule ccontr) case goal1 hence l2:"l = n+1" using l by auto
- hence "\<And>i. i<DIM('a) \<Longrightarrow> \<pi>' i < l" using \<pi>'_i by(auto simp add:less_Suc_eq_le)
- hence *:"\<And> P Q. (\<chi>\<chi> i. if \<pi>' i < l then P i else Q i) = ((\<chi>\<chi> i. P i)::'a)" apply-apply(subst euclidean_eq) by auto
- hence "k = {c..d}" using l(1) \<pi>'_i unfolding * by(auto)
- thus False using `k\<noteq>k'` k' by auto
- qed have **:"\<pi>' (\<pi> l) = l" using \<pi>'_\<pi>[of l] using l ln by auto
- have "x $$ \<pi> l < c $$ \<pi> l \<or> d $$ \<pi> l < x $$ \<pi> l" using l(1) apply-
- proof(erule disjE)
- assume as:"k = ?p1 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
- show ?thesis using *[of "\<pi> l"] using ln l(2) using \<pi>_i[of l] by(auto simp add:** not_less)
- next assume as:"k = ?p2 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
- show ?thesis using *[of "\<pi> l"] using ln l(2) using \<pi>_i[of l] unfolding ** by auto
- qed thus False using x unfolding k' unfolding Int_iff interior_closed_interval mem_interval
- by(auto elim!:allE[where x="\<pi> l"])
- next case False hence "l < n + 1" using l'(2) using `l\<le>l'` by auto
- hence ln:"l \<in> {1..n}" "l' \<in> {1..n}" using l l' False by auto
- note \<pi>_l = \<pi>'_\<pi>[OF ln(1)] \<pi>'_\<pi>[OF ln(2)]
- assume x:"x \<in> interior k \<inter> interior k'"
- show False using l(1) l'(1) apply-
- proof(erule_tac[!] disjE)+
- assume as:"k = ?p1 l" "k' = ?p1 l'"
- note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
- have "l \<noteq> l'" using k'(2)[unfolded as] by auto
- thus False using *[of "\<pi> l'"] *[of "\<pi> l"] ln using \<pi>_i[OF ln(1)] \<pi>_i[OF ln(2)] apply(cases "l<l'")
- by(auto simp add:euclidean_lambda_beta' \<pi>_l \<pi>_i n_def)
- next assume as:"k = ?p2 l" "k' = ?p2 l'"
- note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
- have "l \<noteq> l'" apply(rule) using k'(2)[unfolded as] by auto
- thus False using *[of "\<pi> l"] *[of "\<pi> l'"] `l \<le> l'` ln by(auto simp add:euclidean_lambda_beta' \<pi>_l \<pi>_i n_def)
- next assume as:"k = ?p1 l" "k' = ?p2 l'"
- note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
- show False using abcd[of "\<pi> l'"] using *[of "\<pi> l"] *[of "\<pi> l'"] `l \<le> l'` ln apply(cases "l=l'")
- by(auto simp add:euclidean_lambda_beta' \<pi>_l \<pi>_i n_def)
- next assume as:"k = ?p2 l" "k' = ?p1 l'"
- note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
- show False using *[of "\<pi> l"] *[of "\<pi> l'"] ln `l \<le> l'` apply(cases "l=l'") using abcd[of "\<pi> l'"]
- by(auto simp add:euclidean_lambda_beta' \<pi>_l \<pi>_i n_def)
- qed qed }
- from this[OF k l k' l'] this[OF k'(1) l' k _ l] have "\<And>x. x \<notin> interior k \<inter> interior k'"
- apply - apply(cases "l' \<le> l") using k'(2) by auto
- thus "interior k \<inter> interior k' = {}" by auto
-qed qed
+proof
+ let ?B = "\<lambda>f::'a\<Rightarrow>'a \<times> 'a. {(\<Sum>i\<in>Basis. (fst (f i) \<bullet> i) *\<^sub>R i) .. (\<Sum>i\<in>Basis. (snd (f i) \<bullet> i) *\<^sub>R i)}"
+ def p \<equiv> "?B ` (Basis \<rightarrow>\<^isub>E {(a, c), (c, d), (d, b)})"
+
+ show "{c .. d} \<in> p"
+ unfolding p_def
+ by (auto simp add: interval_eq_empty eucl_le[where 'a='a]
+ intro!: image_eqI[where x="\<lambda>(i::'a)\<in>Basis. (c, d)"])
+
+ { fix i :: 'a assume "i \<in> Basis"
+ with incl nonempty have "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> d \<bullet> i" "d \<bullet> i \<le> b \<bullet> i"
+ unfolding interval_eq_empty subset_interval by (auto simp: not_le) }
+ note ord = this
+
+ show "p division_of {a..b}"
+ proof (rule division_ofI)
+ show "finite p"
+ unfolding p_def by (auto intro!: finite_PiE)
+ { fix k assume "k \<in> p"
+ then obtain f where f: "f \<in> Basis \<rightarrow>\<^isub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f"
+ by (auto simp: p_def)
+ then show "\<exists>a b. k = {a..b}" by auto
+ have "k \<subseteq> {a..b} \<and> k \<noteq> {}"
+ proof (simp add: k interval_eq_empty subset_interval not_less, safe)
+ fix i :: 'a assume i: "i \<in> Basis"
+ moreover with f have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
+ by (auto simp: PiE_iff)
+ moreover note ord[of i]
+ ultimately show "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i"
+ by (auto simp: subset_iff eucl_le[where 'a='a])
+ qed
+ then show "k \<noteq> {}" "k \<subseteq> {a .. b}" by auto
+ {
+ fix l assume "l \<in> p"
+ then obtain g where g: "g \<in> Basis \<rightarrow>\<^isub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g"
+ by (auto simp: p_def)
+ assume "l \<noteq> k"
+ have "\<exists>i\<in>Basis. f i \<noteq> g i"
+ proof (rule ccontr)
+ assume "\<not> (\<exists>i\<in>Basis. f i \<noteq> g i)"
+ with f g have "f = g"
+ by (auto simp: PiE_iff extensional_def intro!: ext)
+ with `l \<noteq> k` show False
+ by (simp add: l k)
+ qed
+ then guess i ..
+ moreover then have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
+ "g i = (a, c) \<or> g i = (c, d) \<or> g i = (d, b)"
+ using f g by (auto simp: PiE_iff)
+ moreover note ord[of i]
+ ultimately show "interior l \<inter> interior k = {}"
+ by (auto simp add: l k interior_closed_interval disjoint_interval intro!: bexI[of _ i]) }
+ note `k \<subseteq> { a.. b}` }
+ moreover
+ { fix x assume x: "x \<in> {a .. b}"
+ have "\<forall>i\<in>Basis. \<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
+ proof
+ fix i :: 'a assume "i \<in> Basis"
+ with x ord[of i]
+ have "(a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> c \<bullet> i) \<or> (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i) \<or>
+ (d \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)"
+ by (auto simp: eucl_le[where 'a='a])
+ then show "\<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
+ by auto
+ qed
+ then guess f unfolding bchoice_iff .. note f = this
+ moreover then have "restrict f Basis \<in> Basis \<rightarrow>\<^isub>E {(a, c), (c, d), (d, b)}"
+ by auto
+ moreover from f have "x \<in> ?B (restrict f Basis)"
+ by (auto simp: mem_interval eucl_le[where 'a='a])
+ ultimately have "\<exists>k\<in>p. x \<in> k"
+ unfolding p_def by blast }
+ ultimately show "\<Union>p = {a..b}"
+ by auto
+ qed
+qed
lemma partial_division_extend_interval: assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> {a..b}"
obtains q where "p \<subseteq> q" "q division_of {a..b::'a::ordered_euclidean_space}" proof(cases "p = {}")
@@ -1415,9 +1297,9 @@
lemma interval_bisection_step: fixes type::"'a::ordered_euclidean_space"
assumes "P {}" "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" "~(P {a..b::'a})"
obtains c d where "~(P{c..d})"
- "\<forall>i<DIM('a). a$$i \<le> c$$i \<and> c$$i \<le> d$$i \<and> d$$i \<le> b$$i \<and> 2 * (d$$i - c$$i) \<le> b$$i - a$$i"
+ "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
proof- have "{a..b} \<noteq> {}" using assms(1,3) by auto
- note ab=this[unfolded interval_eq_empty not_ex not_less]
+ then have ab: "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" by (auto simp: interval_eq_empty not_le)
{ fix f have "finite f \<Longrightarrow>
(\<forall>s\<in>f. P s) \<Longrightarrow>
(\<forall>s\<in>f. \<exists>a b. s = {a..b}) \<Longrightarrow>
@@ -1428,60 +1310,72 @@
apply rule defer apply rule defer apply(rule inter_interior_unions_intervals)
using insert by auto
qed } note * = this
- let ?A = "{{c..d} | c d::'a. \<forall>i<DIM('a). (c$$i = a$$i) \<and> (d$$i = (a$$i + b$$i) / 2) \<or> (c$$i = (a$$i + b$$i) / 2) \<and> (d$$i = b$$i)}"
- let ?PP = "\<lambda>c d. \<forall>i<DIM('a). a$$i \<le> c$$i \<and> c$$i \<le> d$$i \<and> d$$i \<le> b$$i \<and> 2 * (d$$i - c$$i) \<le> b$$i - a$$i"
+ let ?A = "{{c..d} | c d::'a. \<forall>i\<in>Basis. (c\<bullet>i = a\<bullet>i) \<and> (d\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<or> (c\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<and> (d\<bullet>i = b\<bullet>i)}"
+ let ?PP = "\<lambda>c d. \<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
{ presume "\<forall>c d. ?PP c d \<longrightarrow> P {c..d} \<Longrightarrow> False"
thus thesis unfolding atomize_not not_all apply-apply(erule exE)+ apply(rule_tac c=x and d=xa in that) by auto }
assume as:"\<forall>c d. ?PP c d \<longrightarrow> P {c..d}"
have "P (\<Union> ?A)" proof(rule *, rule_tac[2-] ballI, rule_tac[4] ballI, rule_tac[4] impI)
- let ?B = "(\<lambda>s.{(\<chi>\<chi> i. if i \<in> s then a$$i else (a$$i + b$$i) / 2)::'a ..
- (\<chi>\<chi> i. if i \<in> s then (a$$i + b$$i) / 2 else b$$i)}) ` {s. s \<subseteq> {..<DIM('a)}}"
+ let ?B = "(\<lambda>s.{(\<Sum>i\<in>Basis. (if i \<in> s then a\<bullet>i else (a\<bullet>i + b\<bullet>i) / 2) *\<^sub>R i)::'a ..
+ (\<Sum>i\<in>Basis. (if i \<in> s then (a\<bullet>i + b\<bullet>i) / 2 else b\<bullet>i) *\<^sub>R i)}) ` {s. s \<subseteq> Basis}"
have "?A \<subseteq> ?B" proof case goal1
then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+) note c_d=this[rule_format]
have *:"\<And>a b c d. a = c \<Longrightarrow> b = d \<Longrightarrow> {a..b} = {c..d}" by auto
- show "x\<in>?B" unfolding image_iff apply(rule_tac x="{i. i<DIM('a) \<and> c$$i = a$$i}" in bexI)
- unfolding c_d apply(rule * ) unfolding euclidean_eq[where 'a='a] apply safe unfolding euclidean_lambda_beta' mem_Collect_eq
- proof- fix i assume "i<DIM('a)" thus " c $$ i = (if i < DIM('a) \<and> c $$ i = a $$ i then a $$ i else (a $$ i + b $$ i) / 2)"
- "d $$ i = (if i < DIM('a) \<and> c $$ i = a $$ i then (a $$ i + b $$ i) / 2 else b $$ i)"
- using c_d(2)[of i] ab[THEN spec[where x=i]] by(auto simp add:field_simps)
+ show "x\<in>?B" unfolding image_iff
+ apply(rule_tac x="{i. i\<in>Basis \<and> c\<bullet>i = a\<bullet>i}" in bexI)
+ unfolding c_d
+ apply(rule *)
+ apply (simp_all only: euclidean_eq_iff[where 'a='a] inner_setsum_left_Basis mem_Collect_eq simp_thms
+ cong: ball_cong)
+ apply safe
+ proof-
+ fix i :: 'a assume i: "i\<in>Basis"
+ thus " c \<bullet> i = (if c \<bullet> i = a \<bullet> i then a \<bullet> i else (a \<bullet> i + b \<bullet> i) / 2)"
+ "d \<bullet> i = (if c \<bullet> i = a \<bullet> i then (a \<bullet> i + b \<bullet> i) / 2 else b \<bullet> i)"
+ using c_d(2)[of i] ab[OF i] by(auto simp add:field_simps)
qed qed
thus "finite ?A" apply(rule finite_subset) by auto
fix s assume "s\<in>?A" then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+)
note c_d=this[rule_format]
show "P s" unfolding c_d apply(rule as[rule_format]) proof- case goal1 thus ?case
- using c_d(2)[of i] using ab[THEN spec[where x=i]] by auto qed
+ using c_d(2)[of i] using ab[OF `i \<in> Basis`] by auto qed
show "\<exists>a b. s = {a..b}" unfolding c_d by auto
fix t assume "t\<in>?A" then guess e unfolding mem_Collect_eq .. then guess f apply- by(erule exE,(erule conjE)+)
note e_f=this[rule_format]
assume "s \<noteq> t" hence "\<not> (c = e \<and> d = f)" unfolding c_d e_f by auto
- then obtain i where "c$$i \<noteq> e$$i \<or> d$$i \<noteq> f$$i" and i':"i<DIM('a)" unfolding de_Morgan_conj euclidean_eq[where 'a='a] by auto
- hence i:"c$$i \<noteq> e$$i" "d$$i \<noteq> f$$i" apply- apply(erule_tac[!] disjE)
- proof- assume "c$$i \<noteq> e$$i" thus "d$$i \<noteq> f$$i" using c_d(2)[of i] e_f(2)[of i] by fastforce
- next assume "d$$i \<noteq> f$$i" thus "c$$i \<noteq> e$$i" using c_d(2)[of i] e_f(2)[of i] by fastforce
+ then obtain i where "c\<bullet>i \<noteq> e\<bullet>i \<or> d\<bullet>i \<noteq> f\<bullet>i" and i':"i\<in>Basis"
+ unfolding euclidean_eq_iff[where 'a='a] by auto
+ hence i:"c\<bullet>i \<noteq> e\<bullet>i" "d\<bullet>i \<noteq> f\<bullet>i" apply- apply(erule_tac[!] disjE)
+ proof- assume "c\<bullet>i \<noteq> e\<bullet>i" thus "d\<bullet>i \<noteq> f\<bullet>i" using c_d(2)[OF i'] e_f(2)[OF i'] by fastforce
+ next assume "d\<bullet>i \<noteq> f\<bullet>i" thus "c\<bullet>i \<noteq> e\<bullet>i" using c_d(2)[OF i'] e_f(2)[OF i'] by fastforce
qed have *:"\<And>s t. (\<And>a. a\<in>s \<Longrightarrow> a\<in>t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}" by auto
show "interior s \<inter> interior t = {}" unfolding e_f c_d interior_closed_interval proof(rule *)
fix x assume "x\<in>{c<..<d}" "x\<in>{e<..<f}"
- hence x:"c$$i < d$$i" "e$$i < f$$i" "c$$i < f$$i" "e$$i < d$$i" unfolding mem_interval using i'
- apply-apply(erule_tac[!] x=i in allE)+ by auto
+ hence x:"c\<bullet>i < d\<bullet>i" "e\<bullet>i < f\<bullet>i" "c\<bullet>i < f\<bullet>i" "e\<bullet>i < d\<bullet>i" unfolding mem_interval using i'
+ apply-apply(erule_tac[!] x=i in ballE)+ by auto
show False using c_d(2)[OF i'] apply- apply(erule_tac disjE)
- proof(erule_tac[!] conjE) assume as:"c $$ i = a $$ i" "d $$ i = (a $$ i + b $$ i) / 2"
- show False using e_f(2)[of i] and i x unfolding as by(fastforce simp add:field_simps)
- next assume as:"c $$ i = (a $$ i + b $$ i) / 2" "d $$ i = b $$ i"
- show False using e_f(2)[of i] and i x unfolding as by(fastforce simp add:field_simps)
+ proof(erule_tac[!] conjE) assume as:"c \<bullet> i = a \<bullet> i" "d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2"
+ show False using e_f(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps)
+ next assume as:"c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2" "d \<bullet> i = b \<bullet> i"
+ show False using e_f(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps)
qed qed qed
also have "\<Union> ?A = {a..b}" proof(rule set_eqI,rule)
fix x assume "x\<in>\<Union>?A" then guess Y unfolding Union_iff ..
from this(1) guess c unfolding mem_Collect_eq .. then guess d ..
note c_d = this[THEN conjunct2,rule_format] `x\<in>Y`[unfolded this[THEN conjunct1]]
show "x\<in>{a..b}" unfolding mem_interval proof safe
- fix i assume "i<DIM('a)" thus "a $$ i \<le> x $$ i" "x $$ i \<le> b $$ i"
- using c_d(1)[of i] c_d(2)[unfolded mem_interval,THEN spec[where x=i]] by auto qed
+ fix i :: 'a assume i: "i\<in>Basis" thus "a \<bullet> i \<le> x \<bullet> i" "x \<bullet> i \<le> b \<bullet> i"
+ using c_d(1)[OF i] c_d(2)[unfolded mem_interval,THEN bspec, OF i] by auto qed
next fix x assume x:"x\<in>{a..b}"
- have "\<forall>i<DIM('a). \<exists>c d. (c = a$$i \<and> d = (a$$i + b$$i) / 2 \<or> c = (a$$i + b$$i) / 2 \<and> d = b$$i) \<and> c\<le>x$$i \<and> x$$i \<le> d"
- (is "\<forall>i<DIM('a). \<exists>c d. ?P i c d") unfolding mem_interval proof(rule,rule) fix i
- have "?P i (a$$i) ((a $$ i + b $$ i) / 2) \<or> ?P i ((a $$ i + b $$ i) / 2) (b$$i)"
- using x[unfolded mem_interval,THEN spec[where x=i]] by auto thus "\<exists>c d. ?P i c d" by blast
- qed thus "x\<in>\<Union>?A" unfolding Union_iff unfolding lambda_skolem' unfolding Bex_def mem_Collect_eq
+ have "\<forall>i\<in>Basis. \<exists>c d. (c = a\<bullet>i \<and> d = (a\<bullet>i + b\<bullet>i) / 2 \<or> c = (a\<bullet>i + b\<bullet>i) / 2 \<and> d = b\<bullet>i) \<and> c\<le>x\<bullet>i \<and> x\<bullet>i \<le> d"
+ (is "\<forall>i\<in>Basis. \<exists>c d. ?P i c d") unfolding mem_interval
+ proof
+ fix i :: 'a assume i: "i \<in> Basis"
+ have "?P i (a\<bullet>i) ((a \<bullet> i + b \<bullet> i) / 2) \<or> ?P i ((a \<bullet> i + b \<bullet> i) / 2) (b\<bullet>i)"
+ using x[unfolded mem_interval,THEN bspec, OF i] by auto thus "\<exists>c d. ?P i c d" by blast
+ qed
+ thus "x\<in>\<Union>?A"
+ 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 by auto
qed finally show False using assms by auto qed
@@ -1490,8 +1384,8 @@
obtains x where "x \<in> {a..b}" "\<forall>e>0. \<exists>c d. x \<in> {c..d} \<and> {c..d} \<subseteq> ball x e \<and> {c..d} \<subseteq> {a..b} \<and> ~P({c..d})"
proof-
have "\<forall>x. \<exists>y. \<not> P {fst x..snd x} \<longrightarrow> (\<not> P {fst y..snd y} \<and>
- (\<forall>i<DIM('a). fst x$$i \<le> fst y$$i \<and> fst y$$i \<le> snd y$$i \<and> snd y$$i \<le> snd x$$i \<and>
- 2 * (snd y$$i - fst y$$i) \<le> snd x$$i - fst x$$i))" proof case goal1 thus ?case proof-
+ (\<forall>i\<in>Basis. fst x\<bullet>i \<le> fst y\<bullet>i \<and> fst y\<bullet>i \<le> snd y\<bullet>i \<and> snd y\<bullet>i \<le> snd x\<bullet>i \<and>
+ 2 * (snd y\<bullet>i - fst y\<bullet>i) \<le> snd x\<bullet>i - fst x\<bullet>i))" proof case goal1 thus ?case proof-
presume "\<not> P {fst x..snd x} \<Longrightarrow> ?thesis"
thus ?thesis apply(cases "P {fst x..snd x}") by auto
next assume as:"\<not> P {fst x..snd x}" from interval_bisection_step[of P, OF assms(1-2) as] guess c d .
@@ -1499,8 +1393,8 @@
qed qed then guess f apply-apply(drule choice) by(erule exE) note f=this
def AB \<equiv> "\<lambda>n. (f ^^ n) (a,b)" def A \<equiv> "\<lambda>n. fst(AB n)" and B \<equiv> "\<lambda>n. snd(AB n)" note ab_def = this AB_def
have "A 0 = a" "B 0 = b" "\<And>n. \<not> P {A(Suc n)..B(Suc n)} \<and>
- (\<forall>i<DIM('a). A(n)$$i \<le> A(Suc n)$$i \<and> A(Suc n)$$i \<le> B(Suc n)$$i \<and> B(Suc n)$$i \<le> B(n)$$i \<and>
- 2 * (B(Suc n)$$i - A(Suc n)$$i) \<le> B(n)$$i - A(n)$$i)" (is "\<And>n. ?P n")
+ (\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and>
+ 2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>n. ?P n")
proof- show "A 0 = a" "B 0 = b" unfolding ab_def by auto
case goal3 note S = ab_def funpow.simps o_def id_apply show ?case
proof(induct n) case 0 thus ?case unfolding S apply(rule f[rule_format]) using assms(3) by auto
@@ -1508,29 +1402,28 @@
qed qed note AB = this(1-2) conjunctD2[OF this(3),rule_format]
have interv:"\<And>e. 0 < e \<Longrightarrow> \<exists>n. \<forall>x\<in>{A n..B n}. \<forall>y\<in>{A n..B n}. dist x y < e"
- proof- case goal1 guess n using real_arch_pow2[of "(setsum (\<lambda>i. b$$i - a$$i) {..<DIM('a)}) / e"] .. note n=this
+ proof- case goal1 guess n using real_arch_pow2[of "(setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis) / e"] .. note n=this
show ?case apply(rule_tac x=n in exI) proof(rule,rule)
fix x y assume xy:"x\<in>{A n..B n}" "y\<in>{A n..B n}"
- have "dist x y \<le> setsum (\<lambda>i. abs((x - y)$$i)) {..<DIM('a)}" unfolding dist_norm by(rule norm_le_l1)
- also have "\<dots> \<le> setsum (\<lambda>i. B n$$i - A n$$i) {..<DIM('a)}"
- proof(rule setsum_mono) fix i show "\<bar>(x - y) $$ i\<bar> \<le> B n $$ i - A n $$ i"
- using xy[unfolded mem_interval,THEN spec[where x=i]] by auto qed
- also have "\<dots> \<le> setsum (\<lambda>i. b$$i - a$$i) {..<DIM('a)} / 2^n" unfolding setsum_divide_distrib
+ have "dist x y \<le> setsum (\<lambda>i. abs((x - y)\<bullet>i)) Basis" unfolding dist_norm by(rule norm_le_l1)
+ also have "\<dots> \<le> setsum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
+ proof(rule setsum_mono)
+ fix i :: 'a assume i: "i \<in> Basis" show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i"
+ using xy[unfolded mem_interval,THEN bspec, OF i] by (auto simp: inner_diff_left) qed
+ also have "\<dots> \<le> setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n" unfolding setsum_divide_distrib
proof(rule setsum_mono) case goal1 thus ?case
proof(induct n) case 0 thus ?case unfolding AB by auto
- next case (Suc n) have "B (Suc n) $$ i - A (Suc n) $$ i \<le> (B n $$ i - A n $$ i) / 2"
+ next case (Suc n) have "B (Suc n) \<bullet> i - A (Suc n) \<bullet> i \<le> (B n \<bullet> i - A n \<bullet> i) / 2"
using AB(4)[of i n] using goal1 by auto
- also have "\<dots> \<le> (b $$ i - a $$ i) / 2 ^ Suc n" using Suc by(auto simp add:field_simps) finally show ?case .
+ also have "\<dots> \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ Suc n" using Suc by(auto simp add:field_simps) finally show ?case .
qed qed
also have "\<dots> < e" using n using goal1 by(auto simp add:field_simps) finally show "dist x y < e" .
qed qed
- { fix n m ::nat assume "m \<le> n" then guess d unfolding le_Suc_ex_iff .. note d=this
- have "{A n..B n} \<subseteq> {A m..B m}" unfolding d
- proof(induct d) case 0 thus ?case by auto
- next case (Suc d) show ?case apply(rule subset_trans[OF _ Suc])
- apply(rule) unfolding mem_interval apply(rule,erule_tac x=i in allE)
- proof- case goal1 thus ?case using AB(4)[of i "m + d"] by(auto simp add:field_simps)
- qed qed } note ABsubset = this
+ { fix n m :: nat assume "m \<le> n" then have "{A n..B n} \<subseteq> {A m..B m}"
+ proof(induct rule: inc_induct)
+ case (step i) show ?case
+ using AB(4) by (intro order_trans[OF step(2)] subset_interval_imp) auto
+ qed simp } note ABsubset = this
have "\<exists>a. \<forall>n. a\<in>{A n..B n}" apply(rule decreasing_closed_nest[rule_format,OF closed_interval _ ABsubset interv])
proof- fix n show "{A n..B n} \<noteq> {}" apply(cases "0<n") using AB(3)[of "n - 1"] assms(1,3) AB(1-2) by auto qed auto
then guess x0 .. note x0=this[rule_format]
@@ -1784,7 +1677,7 @@
apply(rule integrable_linear) by assumption+
lemma integral_component_eq[simp]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
- assumes "f integrable_on s" shows "integral s (\<lambda>x. f x $$ k) = integral s f $$ k"
+ assumes "f integrable_on s" shows "integral s (\<lambda>x. f x \<bullet> k) = integral s f \<bullet> k"
unfolding integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] ..
lemma has_integral_setsum:
@@ -1852,8 +1745,9 @@
apply(rule integral_unique) using has_integral_empty .
lemma has_integral_refl[intro]: shows "(f has_integral 0) {a..a}" "(f has_integral 0) {a::'a::ordered_euclidean_space}"
-proof- have *:"{a} = {a..a}" apply(rule set_eqI) unfolding mem_interval singleton_iff euclidean_eq[where 'a='a]
- apply safe prefer 3 apply(erule_tac x=i in allE) by(auto simp add: field_simps)
+proof-
+ have *:"{a} = {a..a}" apply(rule set_eqI) unfolding mem_interval singleton_iff euclidean_eq_iff[where 'a='a]
+ apply safe prefer 3 apply(erule_tac x=b in ballE) by(auto simp add: field_simps)
show "(f has_integral 0) {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 by auto qed
@@ -1913,42 +1807,48 @@
subsection {* Additivity of integral on abutting intervals. *}
-lemma interval_split: fixes a::"'a::ordered_euclidean_space" assumes "k<DIM('a)" shows
- "{a..b} \<inter> {x. x$$k \<le> c} = {a .. (\<chi>\<chi> i. if i = k then min (b$$k) c else b$$i)}"
- "{a..b} \<inter> {x. x$$k \<ge> c} = {(\<chi>\<chi> i. if i = k then max (a$$k) c else a$$i) .. b}"
- apply(rule_tac[!] set_eqI) unfolding Int_iff mem_interval mem_Collect_eq using assms by auto
-
-lemma content_split: fixes a::"'a::ordered_euclidean_space" assumes "k<DIM('a)" shows
- "content {a..b} = content({a..b} \<inter> {x. x$$k \<le> c}) + content({a..b} \<inter> {x. x$$k >= c})"
-proof- note simps = interval_split[OF assms] content_closed_interval_cases eucl_le[where 'a='a]
- { presume "a\<le>b \<Longrightarrow> ?thesis" thus ?thesis apply(cases "a\<le>b") unfolding simps using assms by auto }
- have *:"{..<DIM('a)} = insert k ({..<DIM('a)} - {k})" "\<And>x. finite ({..<DIM('a)}-{x})" "\<And>x. x\<notin>{..<DIM('a)}-{x}"
+lemma interval_split:
+ fixes a::"'a::ordered_euclidean_space" assumes "k \<in> Basis"
+ shows
+ "{a..b} \<inter> {x. x\<bullet>k \<le> c} = {a .. (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i) *\<^sub>R i)}"
+ "{a..b} \<inter> {x. x\<bullet>k \<ge> c} = {(\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i) *\<^sub>R i) .. b}"
+ apply(rule_tac[!] set_eqI) unfolding Int_iff mem_interval mem_Collect_eq using assms
+ by auto
+
+lemma content_split: fixes a::"'a::ordered_euclidean_space" assumes "k\<in>Basis" shows
+ "content {a..b} = content({a..b} \<inter> {x. x\<bullet>k \<le> c}) + content({a..b} \<inter> {x. x\<bullet>k >= c})"
+proof cases
+ note simps = interval_split[OF assms] content_closed_interval_cases eucl_le[where 'a='a]
+ have *:"Basis = insert k (Basis - {k})" "\<And>x. finite (Basis-{x})" "\<And>x. x\<notin>Basis-{x}"
using assms by auto
- have *:"\<And>X Y Z. (\<Prod>i\<in>{..<DIM('a)}. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>{..<DIM('a)}-{k}. Z i (Y i))"
- "(\<Prod>i\<in>{..<DIM('a)}. b$$i - a$$i) = (\<Prod>i\<in>{..<DIM('a)}-{k}. b$$i - a$$i) * (b$$k - a$$k)"
+ have *:"\<And>X Y Z. (\<Prod>i\<in>Basis. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>Basis-{k}. Z i (Y i))"
+ "(\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i) = (\<Prod>i\<in>Basis-{k}. b\<bullet>i - a\<bullet>i) * (b\<bullet>k - a\<bullet>k)"
apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto
- assume as:"a\<le>b" moreover have "\<And>x. min (b $$ k) c = max (a $$ k) c
- \<Longrightarrow> x* (b$$k - a$$k) = x*(max (a $$ k) c - a $$ k) + x*(b $$ k - max (a $$ k) c)"
+ assume as:"a\<le>b" moreover have "\<And>x. min (b \<bullet> k) c = max (a \<bullet> k) c
+ \<Longrightarrow> x* (b\<bullet>k - a\<bullet>k) = x*(max (a \<bullet> k) c - a \<bullet> k) + x*(b \<bullet> k - max (a \<bullet> k) c)"
by (auto simp add:field_simps)
- moreover have **:"(\<Prod>i<DIM('a). ((\<chi>\<chi> i. if i = k then min (b $$ k) c else b $$ i)::'a) $$ i - a $$ i) =
- (\<Prod>i<DIM('a). (if i = k then min (b $$ k) c else b $$ i) - a $$ i)"
- "(\<Prod>i<DIM('a). b $$ i - ((\<chi>\<chi> i. if i = k then max (a $$ k) c else a $$ i)::'a) $$ i) =
- (\<Prod>i<DIM('a). b $$ i - (if i = k then max (a $$ k) c else a $$ i))"
- apply(rule_tac[!] setprod.cong) by auto
- have "\<not> a $$ k \<le> c \<Longrightarrow> \<not> c \<le> b $$ k \<Longrightarrow> False"
+ moreover have **:"(\<Prod>i\<in>Basis. ((\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) *\<^sub>R i) \<bullet> i - a \<bullet> i)) =
+ (\<Prod>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) - a \<bullet> i)"
+ "(\<Prod>i\<in>Basis. b \<bullet> i - ((\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i)) =
+ (\<Prod>i\<in>Basis. b \<bullet> i - (if i = k then max (a \<bullet> k) c else a \<bullet> i))"
+ by (auto intro!: setprod_cong)
+ have "\<not> a \<bullet> k \<le> c \<Longrightarrow> \<not> c \<le> b \<bullet> k \<Longrightarrow> False"
unfolding not_le using as[unfolded eucl_le[where 'a='a],rule_format,of k] assms by auto
ultimately show ?thesis using assms unfolding simps **
- unfolding *(1)[of "\<lambda>i x. b$$i - x"] *(1)[of "\<lambda>i x. x - a$$i"] unfolding *(2)
- apply(subst(2) euclidean_lambda_beta''[where 'a='a])
- apply(subst(3) euclidean_lambda_beta''[where 'a='a]) by auto
+ unfolding *(1)[of "\<lambda>i x. b\<bullet>i - x"] *(1)[of "\<lambda>i x. x - a\<bullet>i"] unfolding *(2)
+ by auto
+next
+ assume "\<not> a \<le> b" then have "{a .. b} = {}"
+ unfolding interval_eq_empty by (auto simp: eucl_le[where 'a='a] not_le)
+ then show ?thesis by auto
qed
lemma division_split_left_inj: fixes type::"'a::ordered_euclidean_space"
assumes "d division_of i" "k1 \<in> d" "k2 \<in> d" "k1 \<noteq> k2"
- "k1 \<inter> {x::'a. x$$k \<le> c} = k2 \<inter> {x. x$$k \<le> c}"and k:"k<DIM('a)"
- shows "content(k1 \<inter> {x. x$$k \<le> c}) = 0"
+ "k1 \<inter> {x::'a. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}"and k:"k\<in>Basis"
+ shows "content(k1 \<inter> {x. x\<bullet>k \<le> c}) = 0"
proof- note d=division_ofD[OF assms(1)]
- have *:"\<And>a b::'a. \<And> c. (content({a..b} \<inter> {x. x$$k \<le> c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$$k \<le> c}) = {})"
+ have *:"\<And>a b::'a. \<And> c. (content({a..b} \<inter> {x. x\<bullet>k \<le> c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x\<bullet>k \<le> c}) = {})"
unfolding interval_split[OF k] content_eq_0_interior by auto
guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this
guess u2 v2 using d(4)[OF assms(3)] apply-by(erule exE)+ note uv2=this
@@ -1958,10 +1858,10 @@
lemma division_split_right_inj: fixes type::"'a::ordered_euclidean_space"
assumes "d division_of i" "k1 \<in> d" "k2 \<in> d" "k1 \<noteq> k2"
- "k1 \<inter> {x::'a. x$$k \<ge> c} = k2 \<inter> {x. x$$k \<ge> c}" and k:"k<DIM('a)"
- shows "content(k1 \<inter> {x. x$$k \<ge> c}) = 0"
+ "k1 \<inter> {x::'a. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}" and k:"k\<in>Basis"
+ shows "content(k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0"
proof- note d=division_ofD[OF assms(1)]
- have *:"\<And>a b::'a. \<And> c. (content({a..b} \<inter> {x. x$$k >= c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$$k >= c}) = {})"
+ have *:"\<And>a b::'a. \<And> c. (content({a..b} \<inter> {x. x\<bullet>k >= c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x\<bullet>k >= c}) = {})"
unfolding interval_split[OF k] content_eq_0_interior by auto
guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this
guess u2 v2 using d(4)[OF assms(3)] apply-by(erule exE)+ note uv2=this
@@ -1970,25 +1870,25 @@
defer apply(subst assms(5)[unfolded uv1 uv2]) unfolding uv1 uv2 by auto qed
lemma tagged_division_split_left_inj: fixes x1::"'a::ordered_euclidean_space"
- assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2" "k1 \<inter> {x. x$$k \<le> c} = k2 \<inter> {x. x$$k \<le> c}"
- and k:"k<DIM('a)"
- shows "content(k1 \<inter> {x. x$$k \<le> c}) = 0"
+ assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2" "k1 \<inter> {x. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}"
+ and k:"k\<in>Basis"
+ shows "content(k1 \<inter> {x. x\<bullet>k \<le> c}) = 0"
proof- have *:"\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto
show ?thesis apply(rule division_split_left_inj[OF division_of_tagged_division[OF assms(1)]])
apply(rule_tac[1-2] *) using assms(2-) by auto qed
lemma tagged_division_split_right_inj: fixes x1::"'a::ordered_euclidean_space"
- assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2" "k1 \<inter> {x. x$$k \<ge> c} = k2 \<inter> {x. x$$k \<ge> c}"
- and k:"k<DIM('a)"
- shows "content(k1 \<inter> {x. x$$k \<ge> c}) = 0"
+ assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2" "k1 \<inter> {x. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}"
+ and k:"k\<in>Basis"
+ shows "content(k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0"
proof- have *:"\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto
show ?thesis apply(rule division_split_right_inj[OF division_of_tagged_division[OF assms(1)]])
apply(rule_tac[1-2] *) using assms(2-) by auto qed
lemma division_split: fixes a::"'a::ordered_euclidean_space"
- assumes "p division_of {a..b}" and k:"k<DIM('a)"
- shows "{l \<inter> {x. x$$k \<le> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$$k \<le> c} = {})} division_of({a..b} \<inter> {x. x$$k \<le> c})" (is "?p1 division_of ?I1") and
- "{l \<inter> {x. x$$k \<ge> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$$k \<ge> c} = {})} division_of ({a..b} \<inter> {x. x$$k \<ge> c})" (is "?p2 division_of ?I2")
+ assumes "p division_of {a..b}" and k:"k\<in>Basis"
+ shows "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> p \<and> ~(l \<inter> {x. x\<bullet>k \<le> c} = {})} division_of({a..b} \<inter> {x. x\<bullet>k \<le> c})" (is "?p1 division_of ?I1") and
+ "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> p \<and> ~(l \<inter> {x. x\<bullet>k \<ge> c} = {})} division_of ({a..b} \<inter> {x. x\<bullet>k \<ge> c})" (is "?p2 division_of ?I2")
proof(rule_tac[!] division_ofI) note p=division_ofD[OF assms(1)]
show "finite ?p1" "finite ?p2" using p(1) by auto show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2" unfolding p(6)[THEN sym] by auto
{ fix k assume "k\<in>?p1" then guess l unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l=this
@@ -2006,34 +1906,34 @@
qed
lemma has_integral_split: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector"
- assumes "(f has_integral i) ({a..b} \<inter> {x. x$$k \<le> c})" "(f has_integral j) ({a..b} \<inter> {x. x$$k \<ge> c})" and k:"k<DIM('a)"
+ assumes "(f has_integral i) ({a..b} \<inter> {x. x\<bullet>k \<le> c})" "(f has_integral j) ({a..b} \<inter> {x. x\<bullet>k \<ge> c})" and k:"k\<in>Basis"
shows "(f has_integral (i + j)) ({a..b})"
proof(unfold has_integral,rule,rule) case goal1 hence e:"e/2>0" by auto
guess d1 using has_integralD[OF assms(1)[unfolded interval_split[OF k]] e] . note d1=this[unfolded interval_split[THEN sym,OF k]]
guess d2 using has_integralD[OF assms(2)[unfolded interval_split[OF k]] e] . note d2=this[unfolded interval_split[THEN sym,OF k]]
- let ?d = "\<lambda>x. if x$$k = c then (d1 x \<inter> d2 x) else ball x (abs(x$$k - c)) \<inter> d1 x \<inter> d2 x"
+ let ?d = "\<lambda>x. if x\<bullet>k = c then (d1 x \<inter> d2 x) else ball x (abs(x\<bullet>k - c)) \<inter> d1 x \<inter> d2 x"
show ?case apply(rule_tac x="?d" in exI,rule) defer apply(rule,rule,(erule conjE)+)
proof- 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" note p = this tagged_division_ofD[OF this(1)]
- have lem0:"\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$$k \<le> c} = {}) \<Longrightarrow> x$$k \<le> c"
- "\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$$k \<ge> c} = {}) \<Longrightarrow> x$$k \<ge> c"
+ have lem0:"\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x\<bullet>k \<le> c} = {}) \<Longrightarrow> x\<bullet>k \<le> c"
+ "\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x\<bullet>k \<ge> c} = {}) \<Longrightarrow> x\<bullet>k \<ge> c"
proof- fix x kk assume as:"(x,kk)\<in>p"
- show "~(kk \<inter> {x. x$$k \<le> c} = {}) \<Longrightarrow> x$$k \<le> c"
+ show "~(kk \<inter> {x. x\<bullet>k \<le> c} = {}) \<Longrightarrow> x\<bullet>k \<le> c"
proof(rule ccontr) case goal1
- from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $$ k - c\<bar>"
+ from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>"
using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
- hence "\<exists>y. y \<in> ball x \<bar>x $$ k - c\<bar> \<inter> {x. x $$ k \<le> c}" using goal1(1) by blast
- then guess y .. hence "\<bar>x $$ k - y $$ k\<bar> < \<bar>x $$ k - c\<bar>" "y$$k \<le> c" apply-apply(rule le_less_trans)
- using component_le_norm[of "x - y" k] by(auto simp add:dist_norm)
+ hence "\<exists>y. y \<in> ball x \<bar>x \<bullet> k - c\<bar> \<inter> {x. x \<bullet> k \<le> c}" using goal1(1) by blast
+ then guess y .. hence "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<le> c" apply-apply(rule le_less_trans)
+ using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left)
thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps)
qed
- show "~(kk \<inter> {x. x$$k \<ge> c} = {}) \<Longrightarrow> x$$k \<ge> c"
+ show "~(kk \<inter> {x. x\<bullet>k \<ge> c} = {}) \<Longrightarrow> x\<bullet>k \<ge> c"
proof(rule ccontr) case goal1
- from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $$ k - c\<bar>"
+ from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>"
using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
- hence "\<exists>y. y \<in> ball x \<bar>x $$ k - c\<bar> \<inter> {x. x $$ k \<ge> c}" using goal1(1) by blast
- then guess y .. hence "\<bar>x $$ k - y $$ k\<bar> < \<bar>x $$ k - c\<bar>" "y$$k \<ge> c" apply-apply(rule le_less_trans)
- using component_le_norm[of "x - y" k] by(auto simp add:dist_norm)
+ hence "\<exists>y. y \<in> ball x \<bar>x \<bullet> k - c\<bar> \<inter> {x. x \<bullet> k \<ge> c}" using goal1(1) by blast
+ then guess y .. hence "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<ge> c" apply-apply(rule le_less_trans)
+ using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left)
thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps)
qed
qed
@@ -2053,15 +1953,15 @@
qed auto
have lem4:"\<And>g. (\<lambda>(x,l). content (g l) *\<^sub>R f x) = (\<lambda>(x,l). content l *\<^sub>R f x) o (\<lambda>(x,l). (x,g l))" apply(rule ext) by auto
- let ?M1 = "{(x,kk \<inter> {x. x$$k \<le> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$$k \<le> c} \<noteq> {}}"
+ let ?M1 = "{(x,kk \<inter> {x. x\<bullet>k \<le> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
have "norm ((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) < e/2" apply(rule d1(2),rule tagged_division_ofI)
apply(rule lem2 p(3))+ prefer 6 apply(rule fineI)
- proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = {a..b} \<inter> {x. x$$k \<le> c}" unfolding p(8)[THEN sym] by auto
+ proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = {a..b} \<inter> {x. x\<bullet>k \<le> c}" unfolding p(8)[THEN sym] by auto
fix x l assume xl:"(x,l)\<in>?M1"
then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) . note xl'=this
have "l' \<subseteq> d1 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto
thus "l \<subseteq> d1 x" unfolding xl' by auto
- show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $$ k \<le> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
+ show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x \<bullet> k \<le> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
using lem0(1)[OF xl'(3-4)] by auto
show "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastforce simp add: interval_split[OF k,where c=c])
fix y r let ?goal = "interior l \<inter> interior r = {}" assume yr:"(y,r)\<in>?M1"
@@ -2073,15 +1973,15 @@
thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
qed qed moreover
- let ?M2 = "{(x,kk \<inter> {x. x$$k \<ge> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$$k \<ge> c} \<noteq> {}}"
+ let ?M2 = "{(x,kk \<inter> {x. x\<bullet>k \<ge> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
have "norm ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) < e/2" apply(rule d2(2),rule tagged_division_ofI)
apply(rule lem2 p(3))+ prefer 6 apply(rule fineI)
- proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = {a..b} \<inter> {x. x$$k \<ge> c}" unfolding p(8)[THEN sym] by auto
+ proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = {a..b} \<inter> {x. x\<bullet>k \<ge> c}" unfolding p(8)[THEN sym] by auto
fix x l assume xl:"(x,l)\<in>?M2"
then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) . note xl'=this
have "l' \<subseteq> d2 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto
thus "l \<subseteq> d2 x" unfolding xl' by auto
- show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $$ k \<ge> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
+ show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x \<bullet> k \<ge> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
using lem0(2)[OF xl'(3-4)] by auto
show "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastforce simp add: interval_split[OF k, where c=c])
fix y r let ?goal = "interior l \<inter> interior r = {}" assume yr:"(y,r)\<in>?M2"
@@ -2098,98 +1998,95 @@
also { have *:"\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'b) = 0" using scaleR_zero_left by auto
have "((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)
= (\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - (i + j)" by auto
- also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x $$ k \<le> c}) *\<^sub>R f x) +
- (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x $$ k}) *\<^sub>R f x) - (i + j)"
+ also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) +
+ (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) - (i + j)"
unfolding lem3[OF p(3)] apply(subst setsum_reindex_nonzero[OF p(3)]) defer apply(subst setsum_reindex_nonzero[OF p(3)])
defer unfolding lem4[THEN sym] apply(rule refl) unfolding split_paired_all split_conv apply(rule_tac[!] *)
proof- case goal1 thus ?case apply- apply(rule tagged_division_split_left_inj [OF p(1), of a b aa ba]) using k by auto
next case goal2 thus ?case apply- apply(rule tagged_division_split_right_inj[OF p(1), of a b aa ba]) using k by auto
qed also note setsum_addf[THEN sym]
- also have *:"\<And>x. x\<in>p \<Longrightarrow> (\<lambda>(x, ka). content (ka \<inter> {x. x $$ k \<le> c}) *\<^sub>R f x) x + (\<lambda>(x, ka). content (ka \<inter> {x. c \<le> x $$ k}) *\<^sub>R f x) x
+ also have *:"\<And>x. x\<in>p \<Longrightarrow> (\<lambda>(x, ka). content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) x + (\<lambda>(x, ka). content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) x
= (\<lambda>(x,ka). content ka *\<^sub>R f x) x" unfolding split_paired_all split_conv
proof- fix a b assume "(a,b) \<in> p" from p(6)[OF this] guess u v apply-by(erule exE)+ note uv=this
- thus "content (b \<inter> {x. x $$ k \<le> c}) *\<^sub>R f a + content (b \<inter> {x. c \<le> x $$ k}) *\<^sub>R f a = content b *\<^sub>R f a"
+ thus "content (b \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f a + content (b \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f a = content b *\<^sub>R f a"
unfolding scaleR_left_distrib[THEN sym] unfolding uv content_split[OF k,of u v c] by auto
qed note setsum_cong2[OF this]
- finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x $$ k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x $$ k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i +
- ((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x $$ k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x $$ k} \<noteq> {}}. content k *\<^sub>R f x) - j) =
+ finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x \<bullet> k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i +
+ ((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x \<bullet> k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}}. content k *\<^sub>R f x) - j) =
(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)" by auto }
finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed
-(*lemma has_integral_split_cart: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
- assumes "(f has_integral i) ({a..b} \<inter> {x. x$k \<le> c})" "(f has_integral j) ({a..b} \<inter> {x. x$k \<ge> c})"
- shows "(f has_integral (i + j)) ({a..b})" *)
-
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} \<inter> {x. x$$k \<le> (c::real)})" "p2 tagged_division_of ({a..b} \<inter> {x. x$$k \<ge> c})"
- and k:"k<DIM('a)"
+ assumes "p1 tagged_division_of ({a..b} \<inter> {x. x\<bullet>k \<le> (c::real)})" "p2 tagged_division_of ({a..b} \<inter> {x. x\<bullet>k \<ge> c})"
+ and k:"k\<in>Basis"
shows "(p1 \<union> p2) tagged_division_of ({a..b})"
-proof- have *:"{a..b} = ({a..b} \<inter> {x. x$$k \<le> c}) \<union> ({a..b} \<inter> {x. x$$k \<ge> c})" by auto
+proof- have *:"{a..b} = ({a..b} \<inter> {x. x\<bullet>k \<le> c}) \<union> ({a..b} \<inter> {x. x\<bullet>k \<ge> c})" by auto
show ?thesis apply(subst *) apply(rule tagged_division_union[OF assms(1-2)])
unfolding interval_split[OF k] interior_closed_interval using k
- by(auto simp add: eucl_less[where 'a='a] elim!:allE[where x=k]) qed
+ by(auto simp add: eucl_less[where 'a='a] elim!: ballE[where x=k]) qed
lemma has_integral_separate_sides: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector"
- assumes "(f has_integral i) ({a..b})" "e>0" and k:"k<DIM('a)"
- obtains d where "gauge d" "(\<forall>p1 p2. p1 tagged_division_of ({a..b} \<inter> {x. x$$k \<le> c}) \<and> d fine p1 \<and>
- p2 tagged_division_of ({a..b} \<inter> {x. x$$k \<ge> c}) \<and> d fine p2
+ assumes "(f has_integral i) ({a..b})" "e>0" and k:"k\<in>Basis"
+ obtains d where "gauge d" "(\<forall>p1 p2. p1 tagged_division_of ({a..b} \<inter> {x. x\<bullet>k \<le> c}) \<and> d fine p1 \<and>
+ p2 tagged_division_of ({a..b} \<inter> {x. x\<bullet>k \<ge> c}) \<and> d fine p2
\<longrightarrow> norm((setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 +
setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e)"
proof- guess d using has_integralD[OF assms(1-2)] . note d=this
show ?thesis apply(rule that[of d]) apply(rule d) apply(rule,rule,rule,(erule conjE)+)
- proof- fix p1 p2 assume "p1 tagged_division_of {a..b} \<inter> {x. x $$ k \<le> c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this
- assume "p2 tagged_division_of {a..b} \<inter> {x. c \<le> x $$ k}" "d fine p2" note p2=tagged_division_ofD[OF this(1)] this
+ proof- fix p1 p2 assume "p1 tagged_division_of {a..b} \<inter> {x. x \<bullet> k \<le> c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this
+ assume "p2 tagged_division_of {a..b} \<inter> {x. c \<le> x \<bullet> 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 ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) = norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)"
apply(subst setsum_Un_zero) apply(rule p1 p2)+ apply(rule) unfolding split_paired_all split_conv
proof- fix a b assume ab:"(a,b) \<in> p1 \<inter> p2"
have "(a,b) \<in> p1" using ab by auto from p1(4)[OF this] guess u v apply-by(erule exE)+ note uv =this
- have "b \<subseteq> {x. x$$k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce
- moreover have "interior {x::'a. x $$ k = c} = {}"
- proof(rule ccontr) case goal1 then obtain x where x:"x\<in>interior {x::'a. x$$k = c}" by auto
+ have "b \<subseteq> {x. x\<bullet>k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce
+ moreover have "interior {x::'a. x \<bullet> k = c} = {}"
+ proof(rule ccontr) case goal1 then obtain x where x:"x\<in>interior {x::'a. x\<bullet>k = c}" by auto
then guess e unfolding mem_interior .. note e=this
- have x:"x$$k = c" using x interior_subset by fastforce
- have *:"\<And>i. i<DIM('a) \<Longrightarrow> \<bar>(x - (x + (\<chi>\<chi> i. if i = k then e / 2 else 0))) $$ i\<bar>
- = (if i = k then e/2 else 0)" using e by auto
- have "(\<Sum>i<DIM('a). \<bar>(x - (x + (\<chi>\<chi> i. if i = k then e / 2 else 0))) $$ i\<bar>) =
- (\<Sum>i<DIM('a). (if i = k then e / 2 else 0))" apply(rule setsum_cong2) apply(subst *) by auto
+ have x:"x\<bullet>k = c" using x interior_subset by fastforce
+ have *:"\<And>i. i\<in>Basis \<Longrightarrow> \<bar>(x - (x + (e / 2) *\<^sub>R k)) \<bullet> i\<bar>
+ = (if i = k then e/2 else 0)" using e k by (auto simp: inner_simps inner_not_same_Basis)
+ have "(\<Sum>i\<in>Basis. \<bar>(x - (x + (e / 2 ) *\<^sub>R k)) \<bullet> i\<bar>) =
+ (\<Sum>i\<in>Basis. (if i = k then e / 2 else 0))" apply(rule setsum_cong2) apply(subst *) by auto
also have "... < e" apply(subst setsum_delta) using e by auto
- finally have "x + (\<chi>\<chi> i. if i = k then e/2 else 0) \<in> ball x e" unfolding mem_ball dist_norm
- by(rule le_less_trans[OF norm_le_l1])
- hence "x + (\<chi>\<chi> i. if i = k then e/2 else 0) \<in> {x. x$$k = c}" using e by auto
- thus False unfolding mem_Collect_eq using e x k by auto
+ finally have "x + (e/2) *\<^sub>R k \<in> ball x e"
+ unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1])
+ hence "x + (e/2) *\<^sub>R k \<in> {x. x\<bullet>k = c}" using e by auto
+ thus False unfolding mem_Collect_eq using e x k by (auto simp: inner_simps)
qed ultimately have "content b = 0" unfolding uv content_eq_0_interior apply-apply(drule interior_mono) by auto
thus "content b *\<^sub>R f a = 0" by auto
qed auto
also have "\<dots> < e" by(rule k d(2) p12 fine_union p1 p2)+
finally show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) < e" . qed qed
-lemma integrable_split[intro]: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
- assumes "f integrable_on {a..b}" and k:"k<DIM('a)"
- shows "f integrable_on ({a..b} \<inter> {x. x$$k \<le> c})" (is ?t1) and "f integrable_on ({a..b} \<inter> {x. x$$k \<ge> c})" (is ?t2)
+lemma integrable_split[intro]:
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
+ assumes "f integrable_on {a..b}" and k:"k\<in>Basis"
+ shows "f integrable_on ({a..b} \<inter> {x. x\<bullet>k \<le> c})" (is ?t1) and "f integrable_on ({a..b} \<inter> {x. x\<bullet>k \<ge> c})" (is ?t2)
proof- guess y using assms(1) unfolding integrable_on_def .. note y=this
- def b' \<equiv> "(\<chi>\<chi> i. if i = k then min (b$$k) c else b$$i)::'a"
- and a' \<equiv> "(\<chi>\<chi> i. if i = k then max (a$$k) c else a$$i)::'a"
+ def b' \<equiv> "\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i)*\<^sub>R i::'a"
+ def a' \<equiv> "\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i)*\<^sub>R i::'a"
show ?t1 ?t2 unfolding interval_split[OF k] integrable_cauchy unfolding interval_split[THEN sym,OF k]
proof(rule_tac[!] allI impI)+ fix e::real assume "e>0" hence "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 = "\<lambda>A. \<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of {a..b} \<inter> A \<and> d fine p1
\<and> p2 tagged_division_of {a..b} \<inter> A \<and> d fine p2 \<longrightarrow>
norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e)"
- show "?P {x. x $$ k \<le> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
- proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $$ k \<le> c} \<and> d fine p1
- \<and> p2 tagged_division_of {a..b} \<inter> {x. x $$ k \<le> c} \<and> d fine p2"
+ show "?P {x. x \<bullet> k \<le> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
+ proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x \<bullet> k \<le> c} \<and> d fine p1
+ \<and> p2 tagged_division_of {a..b} \<inter> {x. x \<bullet> k \<le> c} \<and> d fine p2"
show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
proof- guess p using fine_division_exists[OF d(1), of a' b] . note p=this
show ?thesis using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]]
using as unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric]
using p using assms by(auto simp add:algebra_simps)
qed qed
- show "?P {x. x $$ k \<ge> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
- proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $$ k \<ge> c} \<and> d fine p1
- \<and> p2 tagged_division_of {a..b} \<inter> {x. x $$ k \<ge> c} \<and> d fine p2"
+ show "?P {x. x \<bullet> k \<ge> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
+ proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x \<bullet> k \<ge> c} \<and> d fine p1
+ \<and> p2 tagged_division_of {a..b} \<inter> {x. x \<bullet> k \<ge> c} \<and> d fine p2"
show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
proof- guess p using fine_division_exists[OF d(1), of a b'] . note p=this
show ?thesis using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]]
@@ -2203,13 +2100,13 @@
definition operative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (('b::ordered_euclidean_space) set \<Rightarrow> 'a) \<Rightarrow> bool" where
"operative opp f \<equiv>
(\<forall>a b. content {a..b} = 0 \<longrightarrow> f {a..b} = neutral(opp)) \<and>
- (\<forall>a b c. \<forall>k<DIM('b). f({a..b}) =
- opp (f({a..b} \<inter> {x. x$$k \<le> c}))
- (f({a..b} \<inter> {x. x$$k \<ge> c})))"
+ (\<forall>a b c. \<forall>k\<in>Basis. f({a..b}) =
+ opp (f({a..b} \<inter> {x. x\<bullet>k \<le> c}))
+ (f({a..b} \<inter> {x. x\<bullet>k \<ge> c})))"
lemma operativeD[dest]: fixes type::"'a::ordered_euclidean_space" assumes "operative opp f"
shows "\<And>a b. content {a..b} = 0 \<Longrightarrow> f {a..b::'a} = neutral(opp)"
- "\<And>a b c k. k<DIM('a) \<Longrightarrow> f({a..b}) = opp (f({a..b} \<inter> {x. x$$k \<le> c})) (f({a..b} \<inter> {x. x$$k \<ge> c}))"
+ "\<And>a b c k. k\<in>Basis \<Longrightarrow> f({a..b}) = opp (f({a..b} \<inter> {x. x\<bullet>k \<le> c})) (f({a..b} \<inter> {x. x\<bullet>k \<ge> c}))"
using assms unfolding operative_def by auto
lemma operative_trivial:
@@ -2342,18 +2239,20 @@
lemma operative_integral: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::banach"
shows "operative (lifted(op +)) (\<lambda>i. if f integrable_on i then Some(integral i f) else None)"
unfolding operative_def unfolding neutral_lifted[OF monoidal_monoid] neutral_add
- apply(rule,rule,rule,rule) defer apply(rule allI impI)+
-proof- fix a b c k assume k:"k<DIM('a)" show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) =
- lifted op + (if f integrable_on {a..b} \<inter> {x. x $$ k \<le> c} then Some (integral ({a..b} \<inter> {x. x $$ k \<le> c}) f) else None)
- (if f integrable_on {a..b} \<inter> {x. c \<le> x $$ k} then Some (integral ({a..b} \<inter> {x. c \<le> x $$ k}) f) else None)"
+ apply(rule,rule,rule,rule) defer apply(rule allI ballI)+
+proof-
+ fix a b c and k :: 'a assume k:"k\<in>Basis"
+ show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) =
+ lifted op + (if f integrable_on {a..b} \<inter> {x. x \<bullet> k \<le> c} then Some (integral ({a..b} \<inter> {x. x \<bullet> k \<le> c}) f) else None)
+ (if f integrable_on {a..b} \<inter> {x. c \<le> x \<bullet> k} then Some (integral ({a..b} \<inter> {x. c \<le> x \<bullet> k}) f) else None)"
proof(cases "f integrable_on {a..b}")
case True show ?thesis unfolding if_P[OF True] using k apply-
unfolding if_P[OF integrable_split(1)[OF True]] unfolding if_P[OF integrable_split(2)[OF True]]
unfolding lifted.simps option.inject apply(rule integral_unique) apply(rule has_integral_split[OF _ _ k])
apply(rule_tac[!] integrable_integral integrable_split)+ using True k by auto
- next case False have "(\<not> (f integrable_on {a..b} \<inter> {x. x $$ k \<le> c})) \<or> (\<not> ( f integrable_on {a..b} \<inter> {x. c \<le> x $$ k}))"
+ next case False have "(\<not> (f integrable_on {a..b} \<inter> {x. x \<bullet> k \<le> c})) \<or> (\<not> ( f integrable_on {a..b} \<inter> {x. c \<le> x \<bullet> k}))"
proof(rule ccontr) case goal1 hence "f integrable_on {a..b}" apply- unfolding integrable_on_def
- apply(rule_tac x="integral ({a..b} \<inter> {x. x $$ k \<le> c}) f + integral ({a..b} \<inter> {x. x $$ k \<ge> c}) f" in exI)
+ apply(rule_tac x="integral ({a..b} \<inter> {x. x \<bullet> k \<le> c}) f + integral ({a..b} \<inter> {x. x \<bullet> k \<ge> c}) f" in exI)
apply(rule has_integral_split[OF _ _ k]) apply(rule_tac[!] integrable_integral) by auto
thus False using False by auto
qed thus ?thesis using False by auto
@@ -2365,91 +2264,110 @@
subsection {* Points of division of a partition. *}
definition "division_points (k::('a::ordered_euclidean_space) set) d =
- {(j,x). j<DIM('a) \<and> (interval_lowerbound k)$$j < x \<and> x < (interval_upperbound k)$$j \<and>
- (\<exists>i\<in>d. (interval_lowerbound i)$$j = x \<or> (interval_upperbound i)$$j = x)}"
+ {(j,x). j\<in>Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and>
+ (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
lemma division_points_finite: fixes i::"('a::ordered_euclidean_space) set"
assumes "d division_of i" shows "finite (division_points i d)"
proof- note assm = division_ofD[OF assms]
- let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)$$j < x \<and> x < (interval_upperbound i)$$j \<and>
- (\<exists>i\<in>d. (interval_lowerbound i)$$j = x \<or> (interval_upperbound i)$$j = x)}"
- have *:"division_points i d = \<Union>(?M ` {..<DIM('a)})"
+ let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)\<bullet>j < x \<and> x < (interval_upperbound i)\<bullet>j \<and>
+ (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
+ have *:"division_points i d = \<Union>(?M ` Basis)"
unfolding division_points_def by auto
show ?thesis unfolding * using assm by auto qed
lemma division_points_subset: fixes a::"'a::ordered_euclidean_space"
- assumes "d division_of {a..b}" "\<forall>i<DIM('a). a$$i < b$$i" "a$$k < c" "c < b$$k" and k:"k<DIM('a)"
- shows "division_points ({a..b} \<inter> {x. x$$k \<le> c}) {l \<inter> {x. x$$k \<le> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$$k \<le> c} = {})}
+ assumes "d division_of {a..b}" "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" "a\<bullet>k < c" "c < b\<bullet>k" and k:"k\<in>Basis"
+ shows "division_points ({a..b} \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l . l \<in> d \<and> ~(l \<inter> {x. x\<bullet>k \<le> c} = {})}
\<subseteq> division_points ({a..b}) d" (is ?t1) and
- "division_points ({a..b} \<inter> {x. x$$k \<ge> c}) {l \<inter> {x. x$$k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$$k \<ge> c} = {})}
+ "division_points ({a..b} \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x\<bullet>k \<ge> c} = {})}
\<subseteq> division_points ({a..b}) d" (is ?t2)
proof- note assm = division_ofD[OF assms(1)]
- have *:"\<forall>i<DIM('a). a$$i \<le> b$$i" "\<forall>i<DIM('a). a$$i \<le> ((\<chi>\<chi> i. if i = k then min (b $$ k) c else b $$ i)::'a) $$ i"
- "\<forall>i<DIM('a). ((\<chi>\<chi> i. if i = k then max (a $$ k) c else a $$ i)::'a) $$ i \<le> b$$i" "min (b $$ k) c = c" "max (a $$ k) c = c"
+ have *:"\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
+ "\<forall>i\<in>Basis. a\<bullet>i \<le> (\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) *\<^sub>R i) \<bullet> i"
+ "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i \<le> b\<bullet>i"
+ "min (b \<bullet> k) c = c" "max (a \<bullet> k) c = c"
using assms using less_imp_le by auto
- show ?t1 unfolding division_points_def interval_split[OF k, of a b]
- unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding *
- unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+
- unfolding mem_Collect_eq apply(erule exE conjE)+ unfolding euclidean_lambda_beta'
- proof- fix i l x assume as:"a $$ fst x < snd x" "snd x < (if fst x = k then c else b $$ fst x)"
- "interval_lowerbound i $$ fst x = snd x \<or> interval_upperbound i $$ fst x = snd x"
- "i = l \<inter> {x. x $$ k \<le> c}" "l \<in> d" "l \<inter> {x. x $$ k \<le> c} \<noteq> {}" and fstx:"fst x <DIM('a)"
+ show ?t1
+ unfolding division_points_def interval_split[OF k, of a b]
+ unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
+ unfolding *
+ unfolding subset_eq
+ apply(rule)
+ unfolding mem_Collect_eq split_beta
+ apply(erule bexE conjE)+
+ apply(simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms)
+ apply(erule exE conjE)+
+ proof
+ fix i l x assume as:"a \<bullet> fst x < snd x" "snd x < (if fst x = k then c else b \<bullet> fst x)"
+ "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+ "i = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}" and fstx:"fst x \<in>Basis"
from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
- have *:"\<forall>i<DIM('a). u $$ i \<le> ((\<chi>\<chi> i. if i = k then min (v $$ k) c else v $$ i)::'a) $$ i"
+ have *:"\<forall>i\<in>Basis. u \<bullet> i \<le> (\<Sum>i\<in>Basis. (if i = k then min (v \<bullet> k) c else v \<bullet> i) *\<^sub>R i) \<bullet> i"
using as(6) unfolding l interval_split[OF k] interval_ne_empty as .
- have **:"\<forall>i<DIM('a). u$$i \<le> v$$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
- show "fst x <DIM('a) \<and> a $$ fst x < snd x \<and> snd x < b $$ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $$ fst x = snd x
- \<or> interval_upperbound i $$ fst x = snd x)" apply(rule,rule fstx)
- using as(1-3,5) unfolding l interval_split[OF k] interval_ne_empty as interval_bounds[OF *] apply-
- apply(rule,assumption,rule) defer apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **]
- apply(case_tac[!] "fst x = k") using assms fstx apply- unfolding euclidean_lambda_beta by auto
+ have **:"\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
+ show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+ apply (rule bexI[OF _ `l \<in> d`])
+ using as(1-3,5) fstx
+ unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
+ by (auto split: split_if_asm)
+ show "snd x < b \<bullet> fst x"
+ using as(2) `c < b\<bullet>k` by (auto split: split_if_asm)
qed
- show ?t2 unfolding division_points_def interval_split[OF k, of a b]
+ show ?t2
+ unfolding division_points_def interval_split[OF k, of a b]
unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding *
- unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+
- unfolding mem_Collect_eq apply(erule exE conjE)+ unfolding euclidean_lambda_beta' apply(rule,assumption)
- proof- fix i l x assume as:"(if fst x = k then c else a $$ fst x) < snd x" "snd x < b $$ fst x"
- "interval_lowerbound i $$ fst x = snd x \<or> interval_upperbound i $$ fst x = snd x"
- "i = l \<inter> {x. c \<le> x $$ k}" "l \<in> d" "l \<inter> {x. c \<le> x $$ k} \<noteq> {}" and fstx:"fst x < DIM('a)"
+ unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta
+ apply(erule bexE conjE)+
+ apply(simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms)
+ apply(erule exE conjE)+
+ proof
+ fix i l x assume as:"(if fst x = k then c else a \<bullet> fst x) < snd x" "snd x < b \<bullet> fst x"
+ "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+ "i = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}" and fstx:"fst x \<in> Basis"
from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
- have *:"\<forall>i<DIM('a). ((\<chi>\<chi> i. if i = k then max (u $$ k) c else u $$ i)::'a) $$ i \<le> v $$ i"
+ have *:"\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (u \<bullet> k) c else u \<bullet> i) *\<^sub>R i) \<bullet> i \<le> v \<bullet> i"
using as(6) unfolding l interval_split[OF k] interval_ne_empty as .
- have **:"\<forall>i<DIM('a). u$$i \<le> v$$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
- show "a $$ fst x < snd x \<and> snd x < b $$ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $$ fst x = snd x \<or>
- interval_upperbound i $$ fst x = snd x)"
- using as(1-3,5) unfolding l interval_split[OF k] interval_ne_empty as interval_bounds[OF *] apply-
- apply rule defer apply(rule,assumption) apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **]
- apply(case_tac[!] "fst x = k") using assms fstx apply- by(auto simp add:euclidean_lambda_beta'[OF k]) qed qed
+ have **:"\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
+ show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+ apply (rule bexI[OF _ `l \<in> d`])
+ using as(1-3,5) fstx
+ unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
+ by (auto split: split_if_asm)
+ show "a \<bullet> fst x < snd x"
+ using as(1) `a\<bullet>k < c` by (auto split: split_if_asm)
+ qed
+qed
lemma division_points_psubset: fixes a::"'a::ordered_euclidean_space"
- assumes "d division_of {a..b}" "\<forall>i<DIM('a). a$$i < b$$i" "a$$k < c" "c < b$$k"
- "l \<in> d" "interval_lowerbound l$$k = c \<or> interval_upperbound l$$k = c" and k:"k<DIM('a)"
- shows "division_points ({a..b} \<inter> {x. x$$k \<le> c}) {l \<inter> {x. x$$k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x$$k \<le> c} \<noteq> {}}
+ assumes "d division_of {a..b}" "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" "a\<bullet>k < c" "c < b\<bullet>k"
+ "l \<in> d" "interval_lowerbound l\<bullet>k = c \<or> interval_upperbound l\<bullet>k = c" and k:"k\<in>Basis"
+ shows "division_points ({a..b} \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}
\<subset> division_points ({a..b}) d" (is "?D1 \<subset> ?D")
- "division_points ({a..b} \<inter> {x. x$$k \<ge> c}) {l \<inter> {x. x$$k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x$$k \<ge> c} \<noteq> {}}
+ "division_points ({a..b} \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}
\<subset> division_points ({a..b}) d" (is "?D2 \<subset> ?D")
-proof- have ab:"\<forall>i<DIM('a). a$$i \<le> b$$i" using assms(2) by(auto intro!:less_imp_le)
+proof- have ab:"\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i" using assms(2) by(auto intro!:less_imp_le)
guess u v using division_ofD(4)[OF assms(1,5)] apply-by(erule exE)+ note l=this
- have uv:"\<forall>i<DIM('a). u$$i \<le> v$$i" "\<forall>i<DIM('a). a$$i \<le> u$$i \<and> v$$i \<le> b$$i"
+ have uv:"\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "\<forall>i\<in>Basis. a\<bullet>i \<le> u\<bullet>i \<and> v\<bullet>i \<le> b\<bullet>i"
using division_ofD(2,2,3)[OF assms(1,5)] unfolding l interval_ne_empty
unfolding subset_eq apply- defer apply(erule_tac x=u in ballE, erule_tac x=v in ballE) unfolding mem_interval by auto
- have *:"interval_upperbound ({a..b} \<inter> {x. x $$ k \<le> interval_upperbound l $$ k}) $$ k = interval_upperbound l $$ k"
- "interval_upperbound ({a..b} \<inter> {x. x $$ k \<le> interval_lowerbound l $$ k}) $$ k = interval_lowerbound l $$ k"
+ have *:"interval_upperbound ({a..b} \<inter> {x. x \<bullet> k \<le> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
+ "interval_upperbound ({a..b} \<inter> {x. x \<bullet> k \<le> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
unfolding interval_split[OF k] apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab k by auto
have "\<exists>x. x \<in> ?D - ?D1" using assms(2-) apply-apply(erule disjE)
- apply(rule_tac x="(k,(interval_lowerbound l)$$k)" in exI) defer
- apply(rule_tac x="(k,(interval_upperbound l)$$k)" in exI)
+ apply(rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI) defer
+ apply(rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI)
unfolding division_points_def unfolding interval_bounds[OF ab] by(auto simp add:*)
thus "?D1 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) using k by auto
- have *:"interval_lowerbound ({a..b} \<inter> {x. x $$ k \<ge> interval_lowerbound l $$ k}) $$ k = interval_lowerbound l $$ k"
- "interval_lowerbound ({a..b} \<inter> {x. x $$ k \<ge> interval_upperbound l $$ k}) $$ k = interval_upperbound l $$ k"
+ have *:"interval_lowerbound ({a..b} \<inter> {x. x \<bullet> k \<ge> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
+ "interval_lowerbound ({a..b} \<inter> {x. x \<bullet> k \<ge> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
unfolding interval_split[OF k] apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab k by auto
have "\<exists>x. x \<in> ?D - ?D2" using assms(2-) apply-apply(erule disjE)
- apply(rule_tac x="(k,(interval_lowerbound l)$$k)" in exI) defer
- apply(rule_tac x="(k,(interval_upperbound l)$$k)" in exI)
+ apply(rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI) defer
+ apply(rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI)
unfolding division_points_def unfolding interval_bounds[OF ab] by(auto simp add:*)
thus "?D2 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4) k]) by auto qed
@@ -2548,42 +2466,47 @@
using operativeD(1)[OF assms(2)] x by auto
qed qed }
assume "content {a..b} \<noteq> 0" note ab = this[unfolded content_lt_nz[THEN sym] content_pos_lt_eq]
- hence ab':"\<forall>i<DIM('a). a$$i \<le> b$$i" by (auto intro!: less_imp_le) show ?case
+ hence ab':"\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i" by (auto intro!: less_imp_le) show ?case
proof(cases "division_points {a..b} d = {}")
case True have d':"\<forall>i\<in>d. \<exists>u v. i = {u..v} \<and>
- (\<forall>j<DIM('a). u$$j = a$$j \<and> v$$j = a$$j \<or> u$$j = b$$j \<and> v$$j = b$$j \<or> u$$j = a$$j \<and> v$$j = b$$j)"
+ (\<forall>j\<in>Basis. u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j)"
unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule)
- apply(rule_tac x=a in exI,rule_tac x=b in exI) apply(rule,rule refl) apply(rule,rule)
- proof- fix u v j assume j:"j<DIM('a)" assume as:"{u..v} \<in> d" note division_ofD(3)[OF goal1(4) this]
- hence uv:"\<forall>i<DIM('a). u$$i \<le> v$$i" "u$$j \<le> v$$j" using j unfolding interval_ne_empty by auto
- have *:"\<And>p r Q. \<not> j<DIM('a) \<or> p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> (Q {u..v})" using as j by auto
- have "(j, u$$j) \<notin> division_points {a..b} d"
- "(j, v$$j) \<notin> division_points {a..b} d" using True by auto
+ apply(rule_tac x=a in exI,rule_tac x=b in exI) apply(rule,rule refl)
+ proof
+ fix u v and j :: 'a assume j:"j\<in>Basis" assume as:"{u..v} \<in> d" note division_ofD(3)[OF goal1(4) this]
+ hence uv:"\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "u\<bullet>j \<le> v\<bullet>j" using j unfolding interval_ne_empty by auto
+ have *:"\<And>p r Q. \<not> j\<in>Basis \<or> p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> (Q {u..v})" using as j by auto
+ have "(j, u\<bullet>j) \<notin> division_points {a..b} d"
+ "(j, v\<bullet>j) \<notin> division_points {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 have "a$$j \<le> u$$j" "v$$j \<le> b$$j" using division_ofD(2,2,3)[OF goal1(4) as]
+ moreover have "a\<bullet>j \<le> u\<bullet>j" "v\<bullet>j \<le> b\<bullet>j" using division_ofD(2,2,3)[OF goal1(4) as]
unfolding subset_eq apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE)
unfolding interval_ne_empty mem_interval using j by auto
- ultimately show "u$$j = a$$j \<and> v$$j = a$$j \<or> u$$j = b$$j \<and> v$$j = b$$j \<or> u$$j = a$$j \<and> v$$j = b$$j"
+ ultimately show "u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>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) \<in> {a..b}" unfolding mem_interval using ab by(auto intro!:less_imp_le)
+ qed
+ have "(1/2) *\<^sub>R (a+b) \<in> {a..b}"
+ unfolding mem_interval using ab by(auto intro!: less_imp_le simp: inner_simps)
note this[unfolded division_ofD(6)[OF goal1(4),THEN sym] Union_iff]
then guess i .. note i=this guess u v using d'[rule_format,OF i(1)] apply-by(erule exE conjE)+ note uv=this
have "{a..b} \<in> d"
proof- { presume "i = {a..b}" thus ?thesis using i by auto }
{ presume "u = a" "v = b" thus "i = {a..b}" using uv by auto }
- show "u = a" "v = b" unfolding euclidean_eq[where 'a='a]
- proof(safe) fix j assume j:"j<DIM('a)" note i(2)[unfolded uv mem_interval,rule_format,of j]
- thus "u $$ j = a $$ j" "v $$ j = b $$ j" using uv(2)[rule_format,of j] j by auto
+ show "u = a" "v = b" unfolding euclidean_eq_iff[where 'a='a]
+ proof(safe)
+ fix j :: 'a assume j:"j\<in>Basis"
+ note i(2)[unfolded uv mem_interval,rule_format,of j]
+ thus "u \<bullet> j = a \<bullet> j" "v \<bullet> j = b \<bullet> j" using uv(2)[rule_format,of j] j by (auto simp: inner_simps)
qed qed
hence *:"d = insert {a..b} (d - {{a..b}})" by auto
have "iterate opp (d - {{a..b}}) f = neutral opp" apply(rule iterate_eq_neutral[OF goal1(2)])
proof fix x assume x:"x \<in> d - {{a..b}}" hence "x\<in>d" by auto note d'[rule_format,OF this]
then guess u v apply-by(erule exE conjE)+ note uv=this
have "u\<noteq>a \<or> v\<noteq>b" using x[unfolded uv] by auto
- then obtain j where "u$$j \<noteq> a$$j \<or> v$$j \<noteq> b$$j" and j:"j<DIM('a)" unfolding euclidean_eq[where 'a='a] by auto
- hence "u$$j = v$$j" using uv(2)[rule_format,OF j] by auto
- hence "content {u..v} = 0" unfolding content_eq_0 apply(rule_tac x=j in exI) using j by auto
+ then obtain j where "u\<bullet>j \<noteq> a\<bullet>j \<or> v\<bullet>j \<noteq> b\<bullet>j" and j:"j\<in>Basis" unfolding euclidean_eq_iff[where 'a='a] by auto
+ hence "u\<bullet>j = v\<bullet>j" using uv(2)[rule_format,OF j] by auto
+ hence "content {u..v} = 0" unfolding content_eq_0 apply(rule_tac x=j in bexI) using j by auto
thus "f x = neutral opp" unfolding uv(1) by(rule operativeD(1)[OF goal1(3)])
qed thus "iterate opp d f = f {a..b}" apply-apply(subst *)
apply(subst iterate_insert[OF goal1(2)]) using goal1(2,4) by auto
@@ -2591,39 +2514,40 @@
then guess k c unfolding split_paired_Ex apply- unfolding division_points_def mem_Collect_eq split_conv
by(erule exE conjE)+ note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']]
from this(3) guess j .. note j=this
- def d1 \<equiv> "{l \<inter> {x. x$$k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x$$k \<le> c} \<noteq> {}}"
- def d2 \<equiv> "{l \<inter> {x. x$$k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x$$k \<ge> c} \<noteq> {}}"
- def cb \<equiv> "(\<chi>\<chi> i. if i = k then c else b$$i)::'a" and ca \<equiv> "(\<chi>\<chi> i. if i = k then c else a$$i)::'a"
+ def d1 \<equiv> "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
+ def d2 \<equiv> "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
+ def cb \<equiv> "(\<Sum>i\<in>Basis. (if i = k then c else b\<bullet>i) *\<^sub>R i)::'a"
+ def ca \<equiv> "(\<Sum>i\<in>Basis. (if i = k then c else a\<bullet>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)]
- hence *:"(iterate opp d1 f) = f ({a..b} \<inter> {x. x$$k \<le> c})" "(iterate opp d2 f) = f ({a..b} \<inter> {x. x$$k \<ge> c})"
+ hence *:"(iterate opp d1 f) = f ({a..b} \<inter> {x. x\<bullet>k \<le> c})" "(iterate opp d2 f) = f ({a..b} \<inter> {x. x\<bullet>k \<ge> c})"
apply- 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]
unfolding interval_split[OF kc(4)] d1_def[symmetric] d2_def[symmetric] unfolding goal1(2) Suc_le_mono
using goal1(2-3) using division_points_finite[OF goal1(4)] using kc(4) by auto
have "f {a..b} = opp (iterate opp d1 f) (iterate opp d2 f)" (is "_ = ?prev")
unfolding * apply(rule operativeD(2)) using goal1(3) using kc(4) by auto
- also have "iterate opp d1 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$$k \<le> c}))"
+ also have "iterate opp d1 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x\<bullet>k \<le> c}))"
unfolding d1_def apply(rule iterate_nonzero_image_lemma[unfolded o_def])
unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+
unfolding empty_as_interval[THEN sym] apply(rule content_empty)
- proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. x $$ k \<le> c} = y \<inter> {x. x $$ k \<le> c}" "l \<noteq> y"
+ proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} = y \<inter> {x. x \<bullet> k \<le> c}" "l \<noteq> y"
from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this
- show "f (l \<inter> {x. x $$ k \<le> c}) = neutral opp" unfolding l interval_split[OF kc(4)]
+ show "f (l \<inter> {x. x \<bullet> k \<le> c}) = neutral opp" unfolding l interval_split[OF kc(4)]
apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym,OF kc(4)] apply(rule division_split_left_inj)
apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule kc(4) as)+
- qed also have "iterate opp d2 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$$k \<ge> c}))"
+ qed also have "iterate opp d2 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x\<bullet>k \<ge> c}))"
unfolding d2_def apply(rule iterate_nonzero_image_lemma[unfolded o_def])
unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+
unfolding empty_as_interval[THEN sym] apply(rule content_empty)
- proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x $$ k} = y \<inter> {x. c \<le> x $$ k}" "l \<noteq> y"
+ proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} = y \<inter> {x. c \<le> x \<bullet> k}" "l \<noteq> y"
from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this
- show "f (l \<inter> {x. x $$ k \<ge> c}) = neutral opp" unfolding l interval_split[OF kc(4)]
+ show "f (l \<inter> {x. x \<bullet> k \<ge> c}) = neutral opp" unfolding l interval_split[OF kc(4)]
apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym,OF kc(4)] apply(rule division_split_right_inj)
apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as kc(4))+
- qed also have *:"\<forall>x\<in>d. f x = opp (f (x \<inter> {x. x $$ k \<le> c})) (f (x \<inter> {x. c \<le> x $$ k}))"
+ qed also have *:"\<forall>x\<in>d. f x = opp (f (x \<inter> {x. x \<bullet> k \<le> c})) (f (x \<inter> {x. c \<le> x \<bullet> k}))"
unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule,rule operativeD(2)) using goal1(3) kc by auto
- have "opp (iterate opp d (\<lambda>l. f (l \<inter> {x. x $$ k \<le> c}))) (iterate opp d (\<lambda>l. f (l \<inter> {x. c \<le> x $$ k})))
+ have "opp (iterate opp d (\<lambda>l. f (l \<inter> {x. x \<bullet> k \<le> c}))) (iterate opp d (\<lambda>l. f (l \<inter> {x. c \<le> x \<bullet> k})))
= iterate opp d f" apply(subst(3) iterate_eq[OF _ *[rule_format]]) prefer 3
apply(rule iterate_op[THEN sym]) using goal1 by auto
finally show ?thesis by auto
@@ -2754,31 +2678,34 @@
subsection {* Similar theorems about relationship among components. *}
lemma rsum_component_le: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
- assumes "p tagged_division_of {a..b}" "\<forall>x\<in>{a..b}. (f x)$$i \<le> (g x)$$i"
- shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)$$i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)$$i"
- unfolding euclidean_component_setsum apply(rule setsum_mono) apply safe
+ assumes "p tagged_division_of {a..b}" "\<forall>x\<in>{a..b}. (f x)\<bullet>i \<le> (g x)\<bullet>i"
+ shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)\<bullet>i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)\<bullet>i"
+ unfolding inner_setsum_left apply(rule setsum_mono) apply safe
proof- fix a b assume ab:"(a,b) \<in> p" note assm = tagged_division_ofD(2-4)[OF assms(1) ab]
from this(3) guess u v apply-by(erule exE)+ note b=this
- show "(content b *\<^sub>R f a) $$ i \<le> (content b *\<^sub>R g a) $$ i" unfolding b
- unfolding euclidean_simps real_scaleR_def apply(rule mult_left_mono)
+ show "(content b *\<^sub>R f a) \<bullet> i \<le> (content b *\<^sub>R g a) \<bullet> i" unfolding b
+ unfolding inner_simps real_scaleR_def apply(rule mult_left_mono)
defer apply(rule content_pos_le,rule assms(2)[rule_format]) using assm by auto qed
-lemma has_integral_component_le: fixes f g::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
- assumes "(f has_integral i) s" "(g has_integral j) s" "\<forall>x\<in>s. (f x)$$k \<le> (g x)$$k"
- shows "i$$k \<le> j$$k"
+lemma has_integral_component_le:
+ fixes f g::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
+ assumes k: "k \<in> Basis"
+ assumes "(f has_integral i) s" "(g has_integral j) s" "\<forall>x\<in>s. (f x)\<bullet>k \<le> (g x)\<bullet>k"
+ shows "i\<bullet>k \<le> j\<bullet>k"
proof -
have lem:"\<And>a b i (j::'b). \<And>g f::'a \<Rightarrow> 'b. (f has_integral i) ({a..b}) \<Longrightarrow>
- (g has_integral j) ({a..b}) \<Longrightarrow> \<forall>x\<in>{a..b}. (f x)$$k \<le> (g x)$$k \<Longrightarrow> i$$k \<le> j$$k"
+ (g has_integral j) ({a..b}) \<Longrightarrow> \<forall>x\<in>{a..b}. (f x)\<bullet>k \<le> (g x)\<bullet>k \<Longrightarrow> i\<bullet>k \<le> j\<bullet>k"
proof (rule ccontr)
case goal1
- then have *: "0 < (i$$k - j$$k) / 3" by auto
+ then have *: "0 < (i\<bullet>k - j\<bullet>k) / 3" by auto
guess d1 using goal1(1)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d1=this[rule_format]
guess d2 using goal1(2)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d2=this[rule_format]
guess p using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter .
- note p = this(1) conjunctD2[OF this(2)] note le_less_trans[OF component_le_norm, of _ _ k] term g
+ note p = this(1) conjunctD2[OF this(2)]
+ note le_less_trans[OF Basis_le_norm[OF k]]
note this[OF d1(2)[OF conjI[OF p(1,2)]]] this[OF d2(2)[OF conjI[OF p(1,3)]]]
thus False
- unfolding euclidean_simps
+ unfolding inner_simps
using rsum_component_le[OF p(1) goal1(3)]
by (simp add: abs_real_def split: split_if_asm)
qed let ?P = "\<exists>a b. s = {a..b}"
@@ -2787,8 +2714,9 @@
show ?thesis apply(rule lem) using assms[unfolded s] by auto
qed auto } assume as:"\<not> ?P"
{ presume "\<not> ?thesis \<Longrightarrow> False" thus ?thesis by auto }
- assume "\<not> i$$k \<le> j$$k" hence ij:"(i$$k - j$$k) / 3 > 0" by auto
- note has_integral_altD[OF _ as this] from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format]
+ assume "\<not> i\<bullet>k \<le> j\<bullet>k" hence ij:"(i\<bullet>k - j\<bullet>k) / 3 > 0" by auto
+ note has_integral_altD[OF _ as this]
+ from this[OF assms(2)] this[OF assms(3)] guess B1 B2 . note B=this[rule_format]
have "bounded (ball 0 B1 \<union> ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+
from bounded_subset_closed_interval[OF this] guess a b apply- by(erule exE)+
note ab = conjunctD2[OF this[unfolded Un_subset_iff]]
@@ -2796,69 +2724,50 @@
guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this]
have *:"\<And>w1 w2 j i::real .\<bar>w1 - i\<bar> < (i - j) / 3 \<Longrightarrow> \<bar>w2 - j\<bar> < (i - j) / 3 \<Longrightarrow> w1 \<le> w2 \<Longrightarrow> False"
by (simp add: abs_real_def split: split_if_asm)
- note le_less_trans[OF component_le_norm[of _ k]] note this[OF w1(2)] this[OF w2(2)] moreover
- have "w1$$k \<le> w2$$k" apply(rule lem[OF w1(1) w2(1)]) using assms by auto ultimately
- show False unfolding euclidean_simps by(rule *) qed
+ note le_less_trans[OF Basis_le_norm[OF k]] note this[OF w1(2)] this[OF w2(2)] moreover
+ have "w1\<bullet>k \<le> w2\<bullet>k" apply(rule lem[OF w1(1) w2(1)]) using assms by auto ultimately
+ show False unfolding inner_simps by(rule *)
+qed
lemma integral_component_le: fixes g f::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
- assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. (f x)$$k \<le> (g x)$$k"
- shows "(integral s f)$$k \<le> (integral s g)$$k"
+ assumes "k\<in>Basis" "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. (f x)\<bullet>k \<le> (g x)\<bullet>k"
+ shows "(integral s f)\<bullet>k \<le> (integral s g)\<bullet>k"
apply(rule has_integral_component_le) using integrable_integral assms by auto
-(*lemma has_integral_dest_vec1_le: fixes f::"'a::ordered_euclidean_space \<Rightarrow> real^1"
- assumes "(f has_integral i) s" "(g has_integral j) s" "\<forall>x\<in>s. f x \<le> g x"
- shows "dest_vec1 i \<le> dest_vec1 j" apply(rule has_integral_component_le[OF assms(1-2)])
- using assms(3) unfolding vector_le_def by auto
-
-lemma integral_dest_vec1_le: fixes f::"real^'n \<Rightarrow> real^1"
- assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. f x \<le> g x"
- shows "dest_vec1(integral s f) \<le> dest_vec1(integral s g)"
- apply(rule has_integral_dest_vec1_le) apply(rule_tac[1-2] integrable_integral) using assms by auto*)
-
lemma has_integral_component_nonneg: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
- assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> (f x)$$k" shows "0 \<le> i$$k"
- using has_integral_component_le[OF has_integral_0 assms(1)] using assms(2-) by auto
+ assumes "k\<in>Basis" "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> (f x)\<bullet>k" shows "0 \<le> i\<bullet>k"
+ using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] using assms(3-) by auto
lemma integral_component_nonneg: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
- assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> (f x)$$k" shows "0 \<le> (integral s f)$$k"
+ assumes "k\<in>Basis" "f integrable_on s" "\<forall>x\<in>s. 0 \<le> (f x)\<bullet>k" shows "0 \<le> (integral s f)\<bullet>k"
apply(rule has_integral_component_nonneg) using assms by auto
-(*lemma has_integral_dest_vec1_nonneg: fixes f::"real^'n \<Rightarrow> real^1"
- assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> i"
- using has_integral_component_nonneg[OF assms(1), of 1]
- using assms(2) unfolding vector_le_def by auto
-
-lemma integral_dest_vec1_nonneg: fixes f::"real^'n \<Rightarrow> real^1"
- assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> integral s f"
- apply(rule has_integral_dest_vec1_nonneg) using assms by auto*)
-
lemma has_integral_component_neg: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
- assumes "(f has_integral i) s" "\<forall>x\<in>s. (f x)$$k \<le> 0"shows "i$$k \<le> 0"
- using has_integral_component_le[OF assms(1) has_integral_0] assms(2-) by auto
-
-(*lemma has_integral_dest_vec1_neg: fixes f::"real^'n \<Rightarrow> real^1"
- assumes "(f has_integral i) s" "\<forall>x\<in>s. f x \<le> 0" shows "i \<le> 0"
- using has_integral_component_neg[OF assms(1),of 1] using assms(2) by auto*)
-
-lemma has_integral_component_lbound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
- assumes "(f has_integral i) {a..b}" "\<forall>x\<in>{a..b}. B \<le> f(x)$$k" "k<DIM('b)" shows "B * content {a..b} \<le> i$$k"
- using has_integral_component_le[OF has_integral_const assms(1),of "(\<chi>\<chi> i. B)::'b" k] assms(2-)
- unfolding euclidean_simps euclidean_lambda_beta'[OF assms(3)] 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}" "\<forall>x\<in>{a..b}. f x$$k \<le> B" "k<DIM('b)"
- shows "i$$k \<le> B * content({a..b})"
- using has_integral_component_le[OF assms(1) has_integral_const, of k "\<chi>\<chi> i. B"]
- unfolding euclidean_simps euclidean_lambda_beta'[OF assms(3)] using assms(2) by(auto simp add:field_simps)
+ assumes "k\<in>Basis" "(f has_integral i) s" "\<forall>x\<in>s. (f x)\<bullet>k \<le> 0"shows "i\<bullet>k \<le> 0"
+ using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) by auto
+
+lemma has_integral_component_lbound:
+ fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
+ assumes "(f has_integral i) {a..b}" "\<forall>x\<in>{a..b}. B \<le> f(x)\<bullet>k" "k\<in>Basis"
+ shows "B * content {a..b} \<le> i\<bullet>k"
+ using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\<Sum>i\<in>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}" "\<forall>x\<in>{a..b}. f x\<bullet>k \<le> B" "k\<in>Basis"
+ shows "i\<bullet>k \<le> B * content({a..b})"
+ using has_integral_component_le[OF assms(3,1) has_integral_const, of "\<Sum>i\<in>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}" "\<forall>x\<in>{a..b}. B \<le> f(x)$$k" "k<DIM('b)"
- shows "B * content({a..b}) \<le> (integral({a..b}) f)$$k"
+ assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. B \<le> f(x)\<bullet>k" "k\<in>Basis"
+ shows "B * content({a..b}) \<le> (integral({a..b}) f)\<bullet>k"
apply(rule has_integral_component_lbound) using assms unfolding has_integral_integral by auto
lemma integral_component_ubound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
- assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. f(x)$$k \<le> B" "k<DIM('b)"
- shows "(integral({a..b}) f)$$k \<le> B * content({a..b})"
+ assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. f(x)\<bullet>k \<le> B" "k\<in>Basis"
+ shows "(integral({a..b}) f)\<bullet>k \<le> B * content({a..b})"
apply(rule has_integral_component_ubound) using assms unfolding has_integral_integral by auto
subsection {* Uniform limit of integrable functions is integrable. *}
@@ -2949,68 +2858,76 @@
apply(rule iterate_nonzero_image_lemma) apply(rule assms monoidal_monoid)+
unfolding assms using neutral_add unfolding neutral_add using assms by auto
-lemma interval_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "k<DIM('a)"
- shows "{a..b} \<inter> {x . abs(x$$k - c) \<le> (e::real)} =
- {(\<chi>\<chi> i. if i = k then max (a$$k) (c - e) else a$$i) .. (\<chi>\<chi> i. if i = k then min (b$$k) (c + e) else b$$i)}"
+lemma interval_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "k\<in>Basis"
+ shows "{a..b} \<inter> {x . abs(x\<bullet>k - c) \<le> (e::real)} =
+ {(\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) (c - e) else a\<bullet>i) *\<^sub>R i) ..
+ (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) (c + e) else b\<bullet>i) *\<^sub>R i)}"
proof- have *:"\<And>x c e::real. abs(x - c) \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e" by auto
have **:"\<And>s P Q. s \<inter> {x. P x \<and> Q x} = (s \<inter> {x. Q x}) \<inter> {x. P x}" by blast
show ?thesis unfolding * ** interval_split[OF assms] by(rule refl) qed
-lemma division_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "p division_of {a..b}" and k:"k<DIM('a)"
- shows "{l \<inter> {x. abs(x$$k - c) \<le> e} |l. l \<in> p \<and> l \<inter> {x. abs(x$$k - c) \<le> e} \<noteq> {}} division_of ({a..b} \<inter> {x. abs(x$$k - c) \<le> e})"
+lemma division_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "p division_of {a..b}" and k:"k\<in>Basis"
+ shows "{l \<inter> {x. abs(x\<bullet>k - c) \<le> e} |l. l \<in> p \<and> l \<inter> {x. abs(x\<bullet>k - c) \<le> e} \<noteq> {}} division_of ({a..b} \<inter> {x. abs(x\<bullet>k - c) \<le> e})"
proof- have *:"\<And>x c. abs(x - c) \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e" by auto
have **:"\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> p' division_of q'" by auto
note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]]
note division_split(2)[OF this, where c="c-e" and k=k,OF k]
thus ?thesis apply(rule **) using k apply- unfolding interval_doublesplit unfolding * unfolding interval_split interval_doublesplit
apply(rule set_eqI) unfolding mem_Collect_eq apply rule apply(erule conjE exE)+ apply(rule_tac x=la in exI) defer
- apply(erule conjE exE)+ apply(rule_tac x="l \<inter> {x. c + e \<ge> x $$ k}" in exI) apply rule defer apply rule
+ apply(erule conjE exE)+ apply(rule_tac x="l \<inter> {x. c + e \<ge> x \<bullet> k}" in exI) apply rule defer apply rule
apply(rule_tac x=l in exI) by blast+ qed
-lemma content_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "0 < e" and k:"k<DIM('a)"
- obtains d where "0 < d" "content({a..b} \<inter> {x. abs(x$$k - c) \<le> d}) < e"
+lemma content_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "0 < e" and k:"k\<in>Basis"
+ obtains d where "0 < d" "content({a..b} \<inter> {x. abs(x\<bullet>k - c) \<le> d}) < e"
proof(cases "content {a..b} = 0")
case True show ?thesis apply(rule that[of 1]) defer unfolding interval_doublesplit[OF k]
apply(rule le_less_trans[OF content_subset]) defer apply(subst True)
unfolding interval_doublesplit[THEN sym,OF k] using assms by auto
-next case False def d \<equiv> "e / 3 / setprod (\<lambda>i. b$$i - a$$i) ({..<DIM('a)} - {k})"
+next case False def d \<equiv> "e / 3 / setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})"
note False[unfolded content_eq_0 not_ex not_le, rule_format]
- hence "\<And>x. x<DIM('a) \<Longrightarrow> b$$x > a$$x" by(auto simp add:not_le)
- hence prod0:"0 < setprod (\<lambda>i. b$$i - a$$i) ({..<DIM('a)} - {k})" apply-apply(rule setprod_pos) by(auto simp add:field_simps)
+ hence "\<And>x. x\<in>Basis \<Longrightarrow> b\<bullet>x > a\<bullet>x" by(auto simp add:not_le)
+ hence prod0:"0 < setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})" apply-apply(rule setprod_pos) by(auto simp add:field_simps)
hence "d > 0" unfolding d_def using assms by(auto simp add:field_simps) thus ?thesis
- proof(rule that[of d]) have *:"{..<DIM('a)} = insert k ({..<DIM('a)} - {k})" using k by auto
- have **:"{a..b} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d} \<noteq> {} \<Longrightarrow>
- (\<Prod>i\<in>{..<DIM('a)} - {k}. interval_upperbound ({a..b} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) $$ i
- - interval_lowerbound ({a..b} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) $$ i)
- = (\<Prod>i\<in>{..<DIM('a)} - {k}. b$$i - a$$i)" apply(rule setprod_cong,rule refl)
+ proof(rule that[of d]) have *:"Basis = insert k (Basis - {k})" using k by auto
+ have **:"{a..b} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {} \<Longrightarrow>
+ (\<Prod>i\<in>Basis - {k}. interval_upperbound ({a..b} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i
+ - interval_lowerbound ({a..b} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i)
+ = (\<Prod>i\<in>Basis - {k}. b\<bullet>i - a\<bullet>i)" apply(rule setprod_cong,rule refl)
unfolding interval_doublesplit[OF k] apply(subst interval_bounds) defer apply(subst interval_bounds)
unfolding interval_eq_empty not_ex not_less by auto
- show "content ({a..b} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) < e" apply(cases) unfolding content_def apply(subst if_P,assumption,rule assms)
+ show "content ({a..b} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e" apply(cases) unfolding content_def apply(subst if_P,assumption,rule assms)
unfolding if_not_P apply(subst *) apply(subst setprod_insert) unfolding **
unfolding interval_doublesplit[OF k] interval_eq_empty not_ex not_less prefer 3
- apply(subst interval_bounds) defer apply(subst interval_bounds) unfolding euclidean_lambda_beta'[OF k] if_P[OF refl]
- proof- have "(min (b $$ k) (c + d) - max (a $$ k) (c - d)) \<le> 2 * d" by auto
- also have "... < e / (\<Prod>i\<in>{..<DIM('a)} - {k}. b $$ i - a $$ i)" unfolding d_def using assms prod0 by(auto simp add:field_simps)
- finally show "(min (b $$ k) (c + d) - max (a $$ k) (c - d)) * (\<Prod>i\<in>{..<DIM('a)} - {k}. b $$ i - a $$ i) < e"
- unfolding pos_less_divide_eq[OF prod0] . qed auto qed qed
-
-lemma negligible_standard_hyperplane[intro]: fixes type::"'a::ordered_euclidean_space" assumes k:"k<DIM('a)"
- shows "negligible {x::'a. x$$k = (c::real)}"
+ apply(subst interval_bounds) defer apply(subst interval_bounds)
+ apply (simp_all only: k inner_setsum_left_Basis simp_thms if_P cong: bex_cong ball_cong)
+ proof -
+ have "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) \<le> 2 * d" by auto
+ also have "... < e / (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i)" unfolding d_def using assms prod0 by(auto simp add:field_simps)
+ finally show "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) * (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i) < e"
+ unfolding pos_less_divide_eq[OF prod0] .
+ qed auto
+ qed
+qed
+
+lemma negligible_standard_hyperplane[intro]:
+ fixes k :: "'a::ordered_euclidean_space"
+ assumes k: "k \<in> Basis"
+ shows "negligible {x. x\<bullet>k = c}"
unfolding negligible_def has_integral apply(rule,rule,rule,rule)
proof-
case goal1 from content_doublesplit[OF this k,of a b c] guess d . note d=this
- let ?i = "indicator {x::'a. x$$k = c} :: 'a\<Rightarrow>real"
+ let ?i = "indicator {x::'a. x\<bullet>k = c} :: 'a\<Rightarrow>real"
show ?case apply(rule_tac x="\<lambda>x. ball x d" in exI) apply(rule,rule gauge_ball,rule d)
proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \<and> (\<lambda>x. ball x d) fine p"
- have *:"(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. abs(x$$k - c) \<le> d}) *\<^sub>R ?i x)"
+ have *:"(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. abs(x\<bullet>k - c) \<le> d}) *\<^sub>R ?i x)"
apply(rule setsum_cong2) unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv
apply(cases,rule disjI1,assumption,rule disjI2)
- proof- fix x l assume as:"(x,l)\<in>p" "?i x \<noteq> 0" hence xk:"x$$k = c" unfolding indicator_def apply-by(rule ccontr,auto)
- show "content l = content (l \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d})" apply(rule arg_cong[where f=content])
+ proof- fix x l assume as:"(x,l)\<in>p" "?i x \<noteq> 0" hence xk:"x\<bullet>k = c" unfolding indicator_def apply-by(rule ccontr,auto)
+ show "content l = content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})" apply(rule arg_cong[where f=content])
apply(rule set_eqI,rule,rule) unfolding mem_Collect_eq
proof- fix y assume y:"y\<in>l" note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv]
- note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y] note le_less_trans[OF component_le_norm[of _ k] this]
- thus "\<bar>y $$ k - c\<bar> \<le> d" unfolding euclidean_simps xk by auto
+ note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y] note le_less_trans[OF Basis_le_norm[OF k] this]
+ thus "\<bar>y \<bullet> k - c\<bar> \<le> d" unfolding inner_simps xk by auto
qed auto qed
note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]]
show "norm ((\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) - 0) < e" unfolding diff_0_right * unfolding real_scaleR_def real_norm_def
@@ -3018,33 +2935,33 @@
apply(rule mult_nonneg_nonneg) apply(drule p'(4)) apply(erule exE)+ apply(rule_tac b=b in back_subst)
prefer 2 apply(subst(asm) eq_commute) apply assumption
apply(subst interval_doublesplit[OF k]) apply(rule content_pos_le) apply(rule indicator_pos_le)
- proof- have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) * ?i x) \<le> (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}))"
+ proof- have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) \<le> (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}))"
apply(rule setsum_mono) unfolding split_paired_all split_conv
apply(rule mult_right_le_one_le) apply(drule p'(4)) by(auto simp add:interval_doublesplit[OF k])
also have "... < e" apply(subst setsum_over_tagged_division_lemma[OF p[THEN conjunct1]])
- proof- case goal1 have "content ({u..v} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) \<le> content {u..v}"
+ proof- case goal1 have "content ({u..v} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<le> content {u..v}"
unfolding interval_doublesplit[OF k] apply(rule content_subset) unfolding interval_doublesplit[THEN sym,OF k] by auto
thus ?case unfolding goal1 unfolding interval_doublesplit[OF k]
by (blast intro: antisym)
- next have *:"setsum content {l \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d} |l. l \<in> snd ` p \<and> l \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d} \<noteq> {}} \<ge> 0"
+ next have *:"setsum content {l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} |l. l \<in> snd ` p \<and> l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}} \<ge> 0"
apply(rule setsum_nonneg,rule) unfolding mem_Collect_eq image_iff apply(erule exE bexE conjE)+ unfolding split_paired_all
- proof- fix x l a b assume as:"x = l \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}" "(a, b) \<in> p" "l = snd (a, b)"
+ proof- fix x l a b assume as:"x = l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}" "(a, b) \<in> p" "l = snd (a, b)"
guess u v using p'(4)[OF as(2)] apply-by(erule exE)+ note * = this
show "content x \<ge> 0" unfolding as snd_conv * interval_doublesplit[OF k] by(rule content_pos_le)
qed have **:"norm (1::real) \<le> 1" by auto note division_doublesplit[OF p'' k,unfolded interval_doublesplit[OF k]]
note dsum_bound[OF this **,unfolded interval_doublesplit[THEN sym,OF k]]
note this[unfolded real_scaleR_def real_norm_def mult_1_right mult_1, of c d] note le_less_trans[OF this d(2)]
- from this[unfolded abs_of_nonneg[OF *]] show "(\<Sum>ka\<in>snd ` p. content (ka \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d})) < e"
+ from this[unfolded abs_of_nonneg[OF *]] show "(\<Sum>ka\<in>snd ` p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) < e"
apply(subst vsum_nonzero_image_lemma[of "snd ` p" content "{}", unfolded o_def,THEN sym])
apply(rule finite_imageI p' content_empty)+ unfolding forall_in_division[OF p'']
proof(rule,rule,rule,rule,rule,rule,rule,erule conjE) fix m n u v
- assume as:"{m..n} \<in> snd ` p" "{u..v} \<in> snd ` p" "{m..n} \<noteq> {u..v}" "{m..n} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d} = {u..v} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}"
- have "({m..n} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) \<inter> ({u..v} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) \<subseteq> {m..n} \<inter> {u..v}" by blast
+ assume as:"{m..n} \<in> snd ` p" "{u..v} \<in> snd ` p" "{m..n} \<noteq> {u..v}" "{m..n} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} = {u..v} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}"
+ have "({m..n} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<inter> ({u..v} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<subseteq> {m..n} \<inter> {u..v}" by blast
note interior_mono[OF this, unfolded division_ofD(5)[OF p'' as(1-3)] interior_inter[of "{m..n}"]]
- hence "interior ({m..n} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) = {}" unfolding as Int_absorb by auto
- thus "content ({m..n} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) = 0" unfolding interval_doublesplit[OF k] content_eq_0_interior[THEN sym] .
+ hence "interior ({m..n} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = {}" unfolding as Int_absorb by auto
+ thus "content ({m..n} \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = 0" unfolding interval_doublesplit[OF k] content_eq_0_interior[THEN sym] .
qed qed
- finally show "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) * ?i x) < e" .
+ finally show "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) < e" .
qed qed qed
subsection {* A technical lemma about "refinement" of division. *}
@@ -3224,7 +3141,7 @@
using negligible_union by auto
lemma negligible_sing[intro]: "negligible {a::_::ordered_euclidean_space}"
- using negligible_standard_hyperplane[of 0 "a$$0"] by auto
+ using negligible_standard_hyperplane[OF SOME_Basis, of "a \<bullet> (SOME i. i \<in> Basis)"] by auto
lemma negligible_insert[simp]: "negligible(insert a s) \<longleftrightarrow> negligible s"
apply(subst insert_is_Un) unfolding negligible_union_eq by auto
@@ -3276,11 +3193,20 @@
subsection {* In particular, the boundary of an interval is negligible. *}
lemma negligible_frontier_interval: "negligible({a::'a::ordered_euclidean_space..b} - {a<..<b})"
-proof- let ?A = "\<Union>((\<lambda>k. {x. x$$k = a$$k} \<union> {x::'a. x$$k = b$$k}) ` {..<DIM('a)})"
- have "{a..b} - {a<..<b} \<subseteq> ?A" apply rule unfolding Diff_iff mem_interval not_all
- apply(erule conjE exE)+ apply(rule_tac X="{x. x $$ xa = a $$ xa} \<union> {x. x $$ xa = b $$ xa}" in UnionI)
- apply(erule_tac[!] x=xa in allE) by auto
- thus ?thesis apply-apply(rule negligible_subset[of ?A]) apply(rule negligible_unions[OF finite_imageI]) by auto qed
+proof-
+ let ?A = "\<Union>((\<lambda>k. {x. x\<bullet>k = a\<bullet>k} \<union> {x::'a. x\<bullet>k = b\<bullet>k}) ` Basis)"
+ have "{a..b} - {a<..<b} \<subseteq> ?A"
+ apply rule unfolding Diff_iff mem_interval
+ apply simp
+ apply(erule conjE bexE)+
+ apply(rule_tac x=i in bexI)
+ by auto
+ thus ?thesis
+ apply-
+ apply(rule negligible_subset[of ?A])
+ apply(rule negligible_unions[OF finite_imageI])
+ by auto
+qed
lemma has_integral_spike_interior:
assumes "\<forall>x\<in>{a<..<b}. g x = f x" "(f has_integral y) ({a..b})" shows "(g has_integral y) ({a..b})"
@@ -3309,23 +3235,26 @@
lemma operative_approximable: assumes "0 \<le> e" fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::banach"
shows "operative op \<and> (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::'b)) \<le> e) \<and> g integrable_on i)" unfolding operative_def neutral_and
-proof safe fix a b::"'b" { assume "content {a..b} = 0"
+proof safe
+ fix a b::"'b"
+ { assume "content {a..b} = 0"
thus "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}"
apply(rule_tac x=f in exI) using assms by(auto intro!:integrable_on_null) }
- { fix c k g assume as:"\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e" "g integrable_on {a..b}" and k:"k<DIM('b)"
- show "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. x $$ k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. x $$ k \<le> c}"
- "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $$ k}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. c \<le> x $$ k}"
+ { fix c g and k :: 'b
+ assume as:"\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e" "g integrable_on {a..b}" and k:"k\<in>Basis"
+ show "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. x \<bullet> k \<le> c}"
+ "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. c \<le> x \<bullet> k}"
apply(rule_tac[!] x=g in exI) using as(1) integrable_split[OF as(2) k] by auto }
- fix c k g1 g2 assume as:"\<forall>x\<in>{a..b} \<inter> {x. x $$ k \<le> c}. norm (f x - g1 x) \<le> e" "g1 integrable_on {a..b} \<inter> {x. x $$ k \<le> c}"
- "\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $$ k}. norm (f x - g2 x) \<le> e" "g2 integrable_on {a..b} \<inter> {x. c \<le> x $$ k}"
- assume k:"k<DIM('b)"
- let ?g = "\<lambda>x. if x$$k = c then f x else if x$$k \<le> c then g1 x else g2 x"
+ fix c k g1 g2 assume as:"\<forall>x\<in>{a..b} \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g1 x) \<le> e" "g1 integrable_on {a..b} \<inter> {x. x \<bullet> k \<le> c}"
+ "\<forall>x\<in>{a..b} \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g2 x) \<le> e" "g2 integrable_on {a..b} \<inter> {x. c \<le> x \<bullet> k}"
+ assume k:"k\<in>Basis"
+ let ?g = "\<lambda>x. if x\<bullet>k = c then f x else if x\<bullet>k \<le> c then g1 x else g2 x"
show "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}" apply(rule_tac x="?g" in exI)
- proof safe case goal1 thus ?case apply- apply(cases "x$$k=c", case_tac "x$$k < c") using as assms by auto
- next case goal2 presume "?g integrable_on {a..b} \<inter> {x. x $$ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $$ k \<ge> c}"
+ proof safe case goal1 thus ?case apply- apply(cases "x\<bullet>k=c", case_tac "x\<bullet>k < c") using as assms by auto
+ next case goal2 presume "?g integrable_on {a..b} \<inter> {x. x \<bullet> k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x \<bullet> k \<ge> 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} \<inter> {x. x $$ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $$ k \<ge> c}"
+ next show "?g integrable_on {a..b} \<inter> {x. x \<bullet> k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x \<bullet> k \<ge> c}"
apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) using k as(2,4) by auto qed qed
lemma approximable_on_division: fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::banach"
@@ -3355,11 +3284,15 @@
subsection {* Specialization of additivity to one dimension. *}
+lemma
+ shows real_inner_1_left: "inner 1 x = x"
+ and real_inner_1_right: "inner x 1 = x"
+ by simp_all
+
lemma operative_1_lt: assumes "monoidal opp"
shows "operative opp f \<longleftrightarrow> ((\<forall>a b. b \<le> a \<longrightarrow> f {a..b::real} = neutral opp) \<and>
(\<forall>a b c. a < c \<and> c < b \<longrightarrow> opp (f{a..c})(f{c..b}) = f {a..b}))"
- unfolding operative_def content_eq_0 DIM_real less_one simp_thms(39,41) Eucl_real_simps
- (* The dnf_simps simplify "\<exists> x. x= _ \<and> _" and "\<forall>k. k = _ \<longrightarrow> _" *)
+ apply (simp add: operative_def content_eq_0 less_one)
proof safe fix a b c::"real" assume as:"\<forall>a b c. f {a..b} = opp (f ({a..b} \<inter> {x. x \<le> c}))
(f ({a..b} \<inter> {x. c \<le> x}))" "a < c" "c < b"
from this(2-) have "{a..b} \<inter> {x. x \<le> c} = {a..c}" "{a..b} \<inter> {x. x \<ge> c} = {c..b}" by auto
@@ -3376,10 +3309,11 @@
show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto
qed
next case True hence *:"min (b) c = c" "max a c = c" by auto
- have **:"0 < DIM(real)" by auto
- have ***:"\<And>P Q. (\<chi>\<chi> i. if i = 0 then P i else Q i) = (P 0::real)" apply(subst euclidean_eq)
- apply safe unfolding euclidean_lambda_beta' by auto
- show ?thesis unfolding interval_split[OF **,unfolded Eucl_real_simps(1,3)] unfolding *** *
+ have **: "(1::real) \<in> Basis" by simp
+ have ***:"\<And>P Q. (\<Sum>i\<in>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 *** *
proof(cases "c = a \<or> c = b")
case False thus "f {a..b} = opp (f {a..c}) (f {c..b})"
apply-apply(subst as(2)[rule_format]) using True by auto
@@ -3411,13 +3345,13 @@
subsection {* Special case of additivity we need for the FCT. *}
lemma interval_bound_sing[simp]: "interval_upperbound {a} = a" "interval_lowerbound {a} = a"
- unfolding interval_upperbound_def interval_lowerbound_def by auto
+ unfolding interval_upperbound_def interval_lowerbound_def by (auto simp: euclidean_representation)
lemma additive_tagged_division_1: fixes f::"real \<Rightarrow> 'a::real_normed_vector"
assumes "a \<le> b" "p tagged_division_of {a..b}"
shows "setsum (\<lambda>(x,k). f(interval_upperbound k) - f(interval_lowerbound k)) p = f b - f a"
proof- let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
- have ***:"\<forall>i<DIM(real). a $$ i \<le> b $$ i" using assms by auto
+ have ***:"\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i" using assms by auto
have *:"operative op + ?f" unfolding operative_1_lt[OF monoidal_monoid] interval_eq_empty by auto
have **:"{a..b} \<noteq> {}" using assms(1) by auto note operative_tagged_division[OF monoidal_monoid * assms(2)]
note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],THEN sym]
@@ -3518,38 +3452,30 @@
have "k \<subseteq> \<Union>(s - {k})" apply safe apply(rule *[unfolded closed_limpt,rule_format]) unfolding islimpt_approachable
proof safe fix x and e::real assume as:"x\<in>k" "e>0"
from k(2)[unfolded k content_eq_0] guess i ..
- hence i:"c$$i = d$$i" "i<DIM('a)" using s(3)[OF k(1),unfolded k] unfolding interval_ne_empty by auto
- hence xi:"x$$i = d$$i" using as unfolding k mem_interval by (metis antisym)
- def y \<equiv> "(\<chi>\<chi> j. if j = i then if c$$i \<le> (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)::'a"
+ hence i:"c\<bullet>i = d\<bullet>i" "i\<in>Basis" using s(3)[OF k(1),unfolded k] unfolding interval_ne_empty by auto
+ hence xi:"x\<bullet>i = d\<bullet>i" using as unfolding k mem_interval by (metis antisym)
+ def y \<equiv> "(\<Sum>j\<in>Basis. (if j = i then if c\<bullet>i \<le> (a\<bullet>i + b\<bullet>i) / 2 then c\<bullet>i +
+ min e (b\<bullet>i - c\<bullet>i) / 2 else c\<bullet>i - min e (c\<bullet>i - a\<bullet>i) / 2 else x\<bullet>j) *\<^sub>R j)::'a"
show "\<exists>x'\<in>\<Union>(s - {k}). x' \<noteq> x \<and> dist x' x < e" apply(rule_tac x=y in bexI)
proof have "d \<in> {c..d}" using s(3)[OF k(1)] unfolding k interval_eq_empty mem_interval by(fastforce simp add: not_less)
- hence "d \<in> {a..b}" using s(2)[OF k(1)] unfolding k by auto note di = this[unfolded mem_interval,THEN spec[where x=i]]
- hence xyi:"y$$i \<noteq> x$$i" unfolding y_def unfolding i xi euclidean_lambda_beta'[OF i(2)] if_P[OF refl]
- apply(cases) apply(subst if_P,assumption) unfolding if_not_P not_le using as(2)
- using assms(2)[unfolded content_eq_0] using i(2)
- by (auto simp add: not_le min_def)
- thus "y \<noteq> x" unfolding euclidean_eq[where 'a='a] using i by auto
- have *:"{..<DIM('a)} = insert i ({..<DIM('a)} - {i})" using i by auto
- have "norm (y - x) < e + setsum (\<lambda>i. 0) {..<DIM('a)}" apply(rule le_less_trans[OF norm_le_l1])
+ hence "d \<in> {a..b}" using s(2)[OF k(1)] unfolding k by auto note di = this[unfolded mem_interval,THEN bspec[where x=i]]
+ hence xyi:"y\<bullet>i \<noteq> x\<bullet>i"
+ unfolding y_def i xi using as(2) assms(2)[unfolded content_eq_0] i(2)
+ by (auto elim!: ballE[of _ _ i])
+ thus "y \<noteq> x" unfolding euclidean_eq_iff[where 'a='a] using i by auto
+ have *:"Basis = insert i (Basis - {i})" using i by auto
+ have "norm (y - x) < e + setsum (\<lambda>i. 0) Basis" apply(rule le_less_trans[OF norm_le_l1])
apply(subst *,subst setsum_insert) prefer 3 apply(rule add_less_le_mono)
- proof- show "\<bar>(y - x) $$ i\<bar> < e" unfolding y_def euclidean_simps euclidean_lambda_beta'[OF i(2)] if_P[OF refl]
- apply(cases) apply(subst if_P,assumption) unfolding if_not_P unfolding i xi using di as(2) by auto
- show "(\<Sum>i\<in>{..<DIM('a)} - {i}. \<bar>(y - x) $$ i\<bar>) \<le> (\<Sum>i\<in>{..<DIM('a)}. 0)" unfolding y_def by auto
+ proof-
+ show "\<bar>(y - x) \<bullet> i\<bar> < e"
+ using di as(2) y_def i xi by (auto simp: inner_simps)
+ show "(\<Sum>i\<in>Basis - {i}. \<bar>(y - x) \<bullet> i\<bar>) \<le> (\<Sum>i\<in>Basis. 0)"
+ unfolding y_def by (auto simp: inner_simps)
qed auto thus "dist y x < e" unfolding dist_norm by auto
- have "y\<notin>k" unfolding k mem_interval apply rule apply(erule_tac x=i in allE) using xyi unfolding k i xi by auto
- moreover have "y \<in> \<Union>s" unfolding s mem_interval
- proof(rule,rule) note simps = y_def euclidean_lambda_beta' if_not_P
- fix j assume j:"j<DIM('a)" show "a $$ j \<le> y $$ j \<and> y $$ j \<le> b $$ j"
- proof(cases "j = i") case False have "x \<in> {a..b}" using s(2)[OF k(1)] as(1) by auto
- thus ?thesis using j apply- unfolding simps if_not_P[OF False] unfolding mem_interval by auto
- next case True note T = this show ?thesis
- proof(cases "c $$ i \<le> (a $$ i + b $$ i) / 2")
- case True show ?thesis unfolding simps if_P[OF T] if_P[OF True] unfolding i
- using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps)
- next case False thus ?thesis unfolding simps if_P[OF T] if_not_P[OF False] unfolding i
- using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps)
- qed qed qed
+ have "y\<notin>k" unfolding k mem_interval apply rule apply(erule_tac x=i in ballE) using xyi k i xi by auto
+ moreover have "y \<in> \<Union>s"
+ using set_rev_mp[OF as(1) s(2)[OF k(1)]] as(2) di i unfolding s mem_interval y_def
+ by (auto simp: field_simps elim!: ballE[of _ _ i])
ultimately show "y \<in> \<Union>(s - {k})" by auto
qed qed hence "\<Union>(s - {k}) = {a..b}" unfolding s(6)[THEN sym] by auto
hence "{ka \<in> s - {k}. content ka \<noteq> 0} division_of {a..b}" apply-apply(rule assm(2)[rule_format,OF card refl])
@@ -3730,23 +3656,41 @@
"content((\<lambda>x::'a::ordered_euclidean_space. m *\<^sub>R x + c) ` {a..b}) = (abs m) ^ DIM('a) * content {a..b}" (is "?l = ?r")
proof- { presume *:"{a..b}\<noteq>{} \<Longrightarrow> ?thesis" show ?thesis apply(cases,rule *,assumption)
unfolding not_not using content_empty by auto }
- have *:"DIM('a) = card {..<DIM('a)}" by auto
- assume as:"{a..b}\<noteq>{}" show ?thesis proof(cases "m \<ge> 0")
- case True show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_P[OF True]
- unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') defer apply(subst(2) *)
- apply(subst setprod_constant[THEN sym]) apply(rule finite_lessThan) unfolding setprod_timesf[THEN sym]
- apply(rule setprod_cong2) using True as unfolding interval_ne_empty euclidean_simps not_le
- by(auto simp add:field_simps intro:mult_left_mono)
- next case False show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_not_P[OF False]
- unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') defer apply(subst(2) *)
- apply(subst setprod_constant[THEN sym]) apply(rule finite_lessThan) unfolding setprod_timesf[THEN sym]
- apply(rule setprod_cong2) using False as unfolding interval_ne_empty euclidean_simps not_le
- by(auto simp add:field_simps mult_le_cancel_left_neg) qed qed
+ assume as: "{a..b}\<noteq>{}"
+ show ?thesis
+ proof (cases "m \<ge> 0")
+ case True
+ with as have "{m *\<^sub>R a + c..m *\<^sub>R b + c} \<noteq> {}"
+ unfolding interval_ne_empty
+ apply (intro ballI)
+ apply (erule_tac x=i in ballE)
+ apply (auto simp: inner_simps intro!: mult_left_mono)
+ done
+ moreover from True have *: "\<And>i. (m *\<^sub>R b + c) \<bullet> i - (m *\<^sub>R a + c) \<bullet> i = m *\<^sub>R (b - a) \<bullet> i"
+ by (simp add: inner_simps field_simps)
+ ultimately show ?thesis
+ by (simp add: image_affinity_interval True content_closed_interval'
+ setprod_timesf setprod_constant inner_diff_left)
+ next
+ case False
+ moreover with as have "{m *\<^sub>R b + c..m *\<^sub>R a + c} \<noteq> {}"
+ unfolding interval_ne_empty
+ apply (intro ballI)
+ apply (erule_tac x=i in ballE)
+ apply (auto simp: inner_simps intro!: mult_left_mono)
+ done
+ moreover from False have *: "\<And>i. (m *\<^sub>R a + c) \<bullet> i - (m *\<^sub>R b + c) \<bullet> i = (-m) *\<^sub>R (b - a) \<bullet> i"
+ by (simp add: inner_simps field_simps)
+ ultimately show ?thesis
+ by (simp add: image_affinity_interval content_closed_interval'
+ 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}" "m \<noteq> 0"
shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral ((1 / (abs(m) ^ DIM('a))) *\<^sub>R i)) ((\<lambda>x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` {a..b})"
- apply(rule has_integral_twiddle,safe) apply(rule zero_less_power) unfolding euclidean_eq[where 'a='a]
- unfolding scaleR_right_distrib euclidean_simps scaleR_scaleR
+ apply(rule has_integral_twiddle,safe) apply(rule zero_less_power) unfolding euclidean_eq_iff[where 'a='a]
+ unfolding scaleR_right_distrib inner_simps scaleR_scaleR
defer apply(insert assms(2), simp add:field_simps) apply(insert assms(2), simp add:field_simps)
apply(rule continuous_intros)+ apply(rule interval_image_affinity_interval)+ apply(rule content_image_affinity_interval) using assms by auto
@@ -3757,60 +3701,68 @@
subsection {* Special case of stretching coordinate axes separately. *}
lemma image_stretch_interval:
- "(\<lambda>x. \<chi>\<chi> k. m k * x$$k) ` {a..b::'a::ordered_euclidean_space} =
- (if {a..b} = {} then {} else {(\<chi>\<chi> k. min (m(k) * a$$k) (m(k) * b$$k))::'a .. (\<chi>\<chi> k. max (m(k) * a$$k) (m(k) * b$$k))})"
- (is "?l = ?r")
-proof(cases "{a..b}={}") case True thus ?thesis unfolding True by auto
-next have *:"\<And>P Q. (\<forall>i<DIM('a). P i) \<and> (\<forall>i<DIM('a). Q i) \<longleftrightarrow> (\<forall>i<DIM('a). P i \<and> Q i)" by auto
- case False note ab = this[unfolded interval_ne_empty]
- show ?thesis apply-apply(rule set_eqI)
- proof- fix x::"'a" have **:"\<And>P Q. (\<forall>i<DIM('a). P i = Q i) \<Longrightarrow> (\<forall>i<DIM('a). P i) = (\<forall>i<DIM('a). Q i)" by auto
- show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" unfolding if_not_P[OF False]
- unfolding image_iff mem_interval Bex_def euclidean_simps euclidean_eq[where 'a='a] *
- unfolding imp_conjR[THEN sym] apply(subst euclidean_lambda_beta'') apply(subst lambda_skolem'[THEN sym])
- apply(rule **,rule,rule) unfolding euclidean_lambda_beta'
- proof- fix i assume i:"i<DIM('a)" show "(\<exists>xa. (a $$ i \<le> xa \<and> xa \<le> b $$ i) \<and> x $$ i = m i * xa) =
- (min (m i * a $$ i) (m i * b $$ i) \<le> x $$ i \<and> x $$ i \<le> max (m i * a $$ i) (m i * b $$ i))"
- proof(cases "m i = 0") case True thus ?thesis using ab i by auto
- next case False hence "0 < m i \<or> 0 > m i" by auto thus ?thesis apply-
- proof(erule disjE) assume as:"0 < m i" hence *:"min (m i * a $$ i) (m i * b $$ i) = m i * a $$ i"
- "max (m i * a $$ i) (m i * b $$ i) = m i * b $$ i" using ab i unfolding min_def max_def by auto
- show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$$i" in exI)
- using as by(auto simp add:field_simps)
- next assume as:"0 > m i" hence *:"max (m i * a $$ i) (m i * b $$ i) = m i * a $$ i"
- "min (m i * a $$ i) (m i * b $$ i) = m i * b $$ i" using ab as i unfolding min_def max_def
- by(auto simp add:field_simps mult_le_cancel_left_neg intro: order_antisym)
- show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$$i" in exI)
- using as by(auto simp add:field_simps) qed qed qed qed qed
-
-lemma interval_image_stretch_interval: "\<exists>u v. (\<lambda>x. \<chi>\<chi> k. m k * x$$k) ` {a..b::'a::ordered_euclidean_space} = {u..v::'a}"
+ "(\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k)) *\<^sub>R k) ` {a..b::'a::ordered_euclidean_space} =
+ (if {a..b} = {} then {} else
+ {(\<Sum>k\<in>Basis. (min (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k)::'a ..
+ (\<Sum>k\<in>Basis. (max (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k)})"
+proof cases
+ assume *: "{a..b} \<noteq> {}"
+ 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])
+ apply (subst choice_Basis_iff[symmetric])
+ proof (intro allI ball_cong refl)
+ fix x i :: 'a assume "i \<in> Basis"
+ with * have a_le_b: "a \<bullet> i \<le> b \<bullet> i"
+ unfolding interval_ne_empty by auto
+ show "(\<exists>xa. x \<bullet> i = m i * xa \<and> a \<bullet> i \<le> xa \<and> xa \<le> b \<bullet> i) \<longleftrightarrow>
+ min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) \<le> x \<bullet> i \<and> x \<bullet> i \<le> max (m i * (a \<bullet> i)) (m i * (b \<bullet> i))"
+ proof cases
+ assume "m i \<noteq> 0"
+ moreover then have *: "\<And>a b. a = m i * b \<longleftrightarrow> b = a / m i"
+ by (auto simp add: field_simps)
+ moreover have
+ "min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (a \<bullet> i) else m i * (b \<bullet> i))"
+ "max (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (b \<bullet> i) else m i * (a \<bullet> i))"
+ using a_le_b by (auto simp: min_def max_def mult_le_cancel_left)
+ ultimately show ?thesis using a_le_b
+ unfolding * by (auto simp add: le_divide_eq divide_le_eq ac_simps)
+ qed (insert a_le_b, auto)
+ qed
+qed simp
+
+lemma interval_image_stretch_interval:
+ "\<exists>u v. (\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k) ` {a..b::'a::ordered_euclidean_space} = {u..v::'a}"
unfolding image_stretch_interval by auto
lemma content_image_stretch_interval:
- "content((\<lambda>x::'a::ordered_euclidean_space. (\<chi>\<chi> k. m k * x$$k)::'a) ` {a..b}) = abs(setprod m {..<DIM('a)}) * content({a..b})"
+ "content((\<lambda>x::'a::ordered_euclidean_space. (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)::'a) ` {a..b}) = abs(setprod m Basis) * content({a..b})"
proof(cases "{a..b} = {}") case True thus ?thesis
unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto
-next case False hence "(\<lambda>x. (\<chi>\<chi> k. m k * x $$ k)::'a) ` {a..b} \<noteq> {}" by auto
+next case False hence "(\<lambda>x. (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)::'a) ` {a..b} \<noteq> {}" by auto
thus ?thesis using False unfolding content_def image_stretch_interval apply- unfolding interval_bounds' if_not_P
- unfolding abs_setprod setprod_timesf[THEN sym] apply(rule setprod_cong2) unfolding lessThan_iff euclidean_lambda_beta'
- proof- fix i assume i:"i<DIM('a)" have "(m i < 0 \<or> m i > 0) \<or> m i = 0" by auto
- thus "max (m i * a $$ i) (m i * b $$ i) - min (m i * a $$ i) (m i * b $$ i) = \<bar>m i\<bar> * (b $$ i - a $$ i)"
+ unfolding abs_setprod setprod_timesf[THEN sym] apply(rule setprod_cong2) unfolding lessThan_iff
+ proof (simp only: inner_setsum_left_Basis)
+ fix i :: 'a assume i:"i\<in>Basis" have "(m i < 0 \<or> m i > 0) \<or> m i = 0" by auto
+ thus "max (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) - min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) =
+ \<bar>m i\<bar> * (b \<bullet> i - a \<bullet> i)"
apply-apply(erule disjE)+ unfolding min_def max_def using False[unfolded interval_ne_empty,rule_format,of i] i
by(auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos) qed qed
lemma has_integral_stretch: fixes f::"'a::ordered_euclidean_space => 'b::real_normed_vector"
- assumes "(f has_integral i) {a..b}" "\<forall>k<DIM('a). ~(m k = 0)"
- shows "((\<lambda>x. f(\<chi>\<chi> k. m k * x$$k)) has_integral
- ((1/(abs(setprod m {..<DIM('a)}))) *\<^sub>R i)) ((\<lambda>x. (\<chi>\<chi> k. 1/(m k) * x$$k)::'a) ` {a..b})"
+ assumes "(f has_integral i) {a..b}" "\<forall>k\<in>Basis. ~(m k = 0)"
+ shows "((\<lambda>x. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) has_integral
+ ((1/(abs(setprod m Basis))) *\<^sub>R i)) ((\<lambda>x. (\<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k)) ` {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[where 'a='a] using assms
-proof- show "\<forall>y::'a. continuous (at y) (\<lambda>x. (\<chi>\<chi> k. m k * x $$ k)::'a)"
+ unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] using assms
+proof- show "\<forall>y::'a. continuous (at y) (\<lambda>x. (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k))"
apply(rule,rule linear_continuous_at) unfolding linear_linear
- unfolding linear_def euclidean_simps euclidean_eq[where 'a='a] by(auto simp add:field_simps) qed auto
+ unfolding linear_def inner_simps euclidean_eq_iff[where 'a='a] by(auto simp add:field_simps)
+qed auto
lemma integrable_stretch: fixes f::"'a::ordered_euclidean_space => 'b::real_normed_vector"
- assumes "f integrable_on {a..b}" "\<forall>k<DIM('a). ~(m k = 0)"
- shows "(\<lambda>x::'a. f(\<chi>\<chi> k. m k * x$$k)) integrable_on ((\<lambda>x. \<chi>\<chi> k. 1/(m k) * x$$k) ` {a..b})"
+ assumes "f integrable_on {a..b}" "\<forall>k\<in>Basis. ~(m k = 0)"
+ shows "(\<lambda>x::'a. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) integrable_on ((\<lambda>x. \<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k) ` {a..b})"
using assms unfolding integrable_on_def apply-apply(erule exE)
apply(drule has_integral_stretch,assumption) by auto
@@ -3856,7 +3808,8 @@
show ?thesis proof(cases,rule *,assumption)
assume "\<not> a < b" hence "a = b" using assms(1) by auto
hence *:"{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
+ 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 = "\<lambda>e. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow>
norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content {a..b})"
@@ -4106,11 +4059,14 @@
from this(2)[unfolded open_contains_ball,rule_format,OF this(1)] guess k .. note k=conjunctD2[OF this]
let ?d = "min k (c - a)/2" show ?thesis apply(rule that[of ?d])
- proof safe show "?d > 0" using k(1) using assms(2) by auto
- fix t assume as:"c - ?d < t" "t \<le> c" let ?thesis = "norm (integral {a..c} f - integral {a..t} f) < e"
+ proof safe
+ show "?d > 0" using k(1) using assms(2) by auto
+ fix t assume as:"c - ?d < t" "t \<le> c"
+ let ?thesis = "norm (integral {a..c} f - integral {a..t} f) < e"
{ presume *:"t < c \<Longrightarrow> ?thesis"
show ?thesis apply(cases "t = c") defer apply(rule *)
- apply(subst less_le) using `e>0` as(2) by auto } assume "t < c"
+ apply(subst less_le) using `e>0` as(2) by auto }
+ assume "t < c"
have "f integrable_on {a..t}" apply(rule integrable_subinterval[OF assms(1)]) using assms(2-3) as(2) by auto
from integrable_integral[OF this,unfolded has_integral,rule_format,OF *] guess d2 ..
@@ -4123,10 +4079,10 @@
with p(2) have "d2 fine p" unfolding fine_def d3_def apply safe apply(erule_tac x="(a,b)" in ballE)+ by auto
note d2_fin = d2(2)[OF conjI[OF p(1) this]]
- have *:"{a..c} \<inter> {x. x $$0 \<le> t} = {a..t}" "{a..c} \<inter> {x. x$$0 \<ge> t} = {t..c}"
+ have *:"{a..c} \<inter> {x. x \<bullet> 1 \<le> t} = {a..t}" "{a..c} \<inter> {x. x \<bullet> 1 \<ge> t} = {t..c}"
using assms(2-3) as by(auto simp add:field_simps)
have "p \<union> {(c, {t..c})} tagged_division_of {a..c} \<and> d1 fine p \<union> {(c, {t..c})}" apply rule
- apply(rule tagged_division_union_interval[of _ _ _ 0 "t"]) unfolding * apply(rule p)
+ apply(rule tagged_division_union_interval[of _ _ _ 1 "t"]) unfolding * apply(rule p)
apply(rule tagged_division_of_self) unfolding fine_def
proof safe fix x k y assume "(x,k)\<in>p" "y\<in>k" thus "y\<in>d1 x"
using p(2) pt unfolding fine_def d3_def apply- apply(erule_tac x="(x,k)" in ballE)+ by auto
@@ -4178,7 +4134,7 @@
{ presume *:"a<b \<Longrightarrow> ?thesis"
show ?thesis apply(cases,rule *,assumption)
proof- case goal1 hence "{a..b} = {x}" using as(1) apply-apply(rule set_eqI)
- unfolding atLeastAtMost_iff by(auto simp only:field_simps not_less DIM_real)
+ unfolding atLeastAtMost_iff by(auto simp only:field_simps not_less)
thus ?case using `e>0` by auto
qed } assume "a<b"
have "(x=a \<or> x=b) \<or> (a<x \<and> x<b)" using as(1) by (auto simp add:)
@@ -4355,65 +4311,41 @@
qed(insert B `e>0`, auto)
next assume as:"\<forall>e>0. ?r e"
from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format]
- def c \<equiv> "(\<chi>\<chi> i. - max B C)::'n::ordered_euclidean_space" and d \<equiv> "(\<chi>\<chi> i. max B C)::'n::ordered_euclidean_space"
+ def c \<equiv> "(\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)::'n::ordered_euclidean_space"
+ def d \<equiv> "(\<Sum>i\<in>Basis. max B C *\<^sub>R i)::'n::ordered_euclidean_space"
have c_d:"{a..b} \<subseteq> {c..d}" apply safe apply(drule B(2)) unfolding mem_interval
- proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def
- by(auto simp add:field_simps) qed
+ proof
+ case goal1 thus ?case using Basis_le_norm[OF `i\<in>Basis`, of x] unfolding c_def d_def
+ by(auto simp add:field_simps setsum_negf)
+ qed
have "ball 0 C \<subseteq> {c..d}" apply safe unfolding mem_interval mem_ball dist_norm
- proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def by auto qed
+ proof
+ case goal1 thus ?case
+ using Basis_le_norm[OF `i\<in>Basis`, of x] unfolding c_def d_def by (auto simp: setsum_negf)
+ qed
from C(2)[OF this] have "\<exists>y. (f has_integral y) {a..b}"
unfolding has_integral_restrict_closed_subintervals_eq[OF c_d,THEN sym] unfolding s by auto
then guess y .. note y=this
have "y = i" proof(rule ccontr) assume "y\<noteq>i" hence "0 < norm (y - i)" by auto
from as[rule_format,OF this] guess C .. note C=conjunctD2[OF this,rule_format]
- def c \<equiv> "(\<chi>\<chi> i. - max B C)::'n::ordered_euclidean_space" and d \<equiv> "(\<chi>\<chi> i. max B C)::'n::ordered_euclidean_space"
+ def c \<equiv> "(\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)::'n::ordered_euclidean_space"
+ def d \<equiv> "(\<Sum>i\<in>Basis. max B C *\<^sub>R i)::'n::ordered_euclidean_space"
have c_d:"{a..b} \<subseteq> {c..d}" apply safe apply(drule B(2)) unfolding mem_interval
- proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def
- by(auto simp add:field_simps) qed
+ proof case goal1 thus ?case using Basis_le_norm[of i x] unfolding c_def d_def
+ by(auto simp add:field_simps setsum_negf) qed
have "ball 0 C \<subseteq> {c..d}" apply safe unfolding mem_interval mem_ball dist_norm
- proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def by auto qed
+ proof case goal1 thus ?case using Basis_le_norm[of i x] unfolding c_def d_def by (auto simp: setsum_negf) qed
note C(2)[OF this] then guess z .. note z = conjunctD2[OF this, unfolded s]
note this[unfolded has_integral_restrict_closed_subintervals_eq[OF c_d]]
hence "z = y" "norm (z - i) < norm (y - i)" apply- apply(rule has_integral_unique[OF _ y(1)]) .
thus False by auto qed
thus ?l using y unfolding s by auto qed qed
-(*lemma has_integral_trans[simp]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real" shows
- "((\<lambda>x. vec1 (f x)) has_integral vec1 i) s \<longleftrightarrow> (f has_integral i) s"
- unfolding has_integral'[unfolded has_integral]
-proof case goal1 thus ?case apply safe
- apply(erule_tac x=e in allE,safe) apply(rule_tac x=B in exI,safe)
- apply(erule_tac x=a in allE, erule_tac x=b in allE,safe)
- apply(rule_tac x="dest_vec1 z" in exI,safe) apply(erule_tac x=ea in allE,safe)
- apply(rule_tac x=d in exI,safe) apply(erule_tac x=p in allE,safe)
- apply(subst(asm)(2) norm_vector_1) unfolding split_def
- unfolding setsum_component Cart_nth.diff cond_value_iff[of dest_vec1]
- Cart_nth.scaleR vec1_dest_vec1 zero_index apply assumption
- apply(subst(asm)(2) norm_vector_1) by auto
-next case goal2 thus ?case apply safe
- apply(erule_tac x=e in allE,safe) apply(rule_tac x=B in exI,safe)
- apply(erule_tac x=a in allE, erule_tac x=b in allE,safe)
- apply(rule_tac x="vec1 z" in exI,safe) apply(erule_tac x=ea in allE,safe)
- apply(rule_tac x=d in exI,safe) apply(erule_tac x=p in allE,safe)
- apply(subst norm_vector_1) unfolding split_def
- unfolding setsum_component Cart_nth.diff cond_value_iff[of dest_vec1]
- Cart_nth.scaleR vec1_dest_vec1 zero_index apply assumption
- apply(subst norm_vector_1) by auto qed
-
-lemma integral_trans[simp]: assumes "(f::'n::ordered_euclidean_space \<Rightarrow> real) integrable_on s"
- shows "integral s (\<lambda>x. vec1 (f x)) = vec1 (integral s f)"
- apply(rule integral_unique) using assms by auto
-
-lemma integrable_on_trans[simp]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real" shows
- "(\<lambda>x. vec1 (f x)) integrable_on s \<longleftrightarrow> (f integrable_on s)"
- unfolding integrable_on_def
- apply(subst(2) vec1_dest_vec1(1)[THEN sym]) unfolding has_integral_trans
- apply safe defer apply(rule_tac x="vec1 y" in exI) by auto *)
-
lemma has_integral_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
assumes "(f has_integral i) s" "(g has_integral j) s" "\<forall>x\<in>s. (f x) \<le> (g x)"
- shows "i \<le> j" using has_integral_component_le[OF assms(1-2), of 0] using assms(3) by auto
+ shows "i \<le> j"
+ using has_integral_component_le[OF _ assms(1-2), of 1] using assms(3) by auto
lemma integral_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. f x \<le> g x"
@@ -4422,7 +4354,7 @@
lemma has_integral_nonneg: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> i"
- using has_integral_component_nonneg[of "f" "i" s 0]
+ using has_integral_component_nonneg[of 1 f i s]
unfolding o_def using assms by auto
lemma integral_nonneg: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
@@ -4519,19 +4451,20 @@
subsection {* More lemmas that are useful later. *}
lemma has_integral_subset_component_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
- assumes "s \<subseteq> t" "(f has_integral i) s" "(f has_integral j) t" "\<forall>x\<in>t. 0 \<le> f(x)$$k"
- shows "i$$k \<le> j$$k"
+ assumes k: "k\<in>Basis" and as: "s \<subseteq> t" "(f has_integral i) s" "(f has_integral j) t" "\<forall>x\<in>t. 0 \<le> f(x)\<bullet>k"
+ shows "i\<bullet>k \<le> j\<bullet>k"
proof- note has_integral_restrict_univ[THEN sym, of f]
- note assms(2-3)[unfolded this] note * = has_integral_component_le[OF this]
- show ?thesis apply(rule *) using assms(1,4) by auto qed
+ note as(2-3)[unfolded this] note * = has_integral_component_le[OF k this]
+ show ?thesis apply(rule *) using as(1,4) by auto qed
lemma has_integral_subset_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
- assumes "s \<subseteq> t" "(f has_integral i) s" "(f has_integral j) t" "\<forall>x\<in>t. 0 \<le> f(x)"
- shows "i \<le> j" using has_integral_subset_component_le[OF assms(1), of "f" "i" "j" 0] using assms by auto
+ assumes as: "s \<subseteq> t" "(f has_integral i) s" "(f has_integral j) t" "\<forall>x\<in>t. 0 \<le> f(x)"
+ shows "i \<le> j"
+ using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] using assms by auto
lemma integral_subset_component_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
- assumes "s \<subseteq> t" "f integrable_on s" "f integrable_on t" "\<forall>x \<in> t. 0 \<le> f(x)$$k"
- shows "(integral s f)$$k \<le> (integral t f)$$k"
+ assumes "k\<in>Basis" "s \<subseteq> t" "f integrable_on s" "f integrable_on t" "\<forall>x \<in> t. 0 \<le> f(x)\<bullet>k"
+ shows "(integral s f)\<bullet>k \<le> (integral t f)\<bullet>k"
apply(rule has_integral_subset_component_le) using assms by auto
lemma integral_subset_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
@@ -4553,13 +4486,13 @@
let ?f = "\<lambda>x. if x \<in> s then f x else 0"
show ?r proof safe fix a b::"'n::ordered_euclidean_space"
from as[OF zero_less_one] guess B .. note B=conjunctD2[OF this,rule_format]
- let ?a = "(\<chi>\<chi> i. min (a$$i) (-B))::'n::ordered_euclidean_space" and ?b = "(\<chi>\<chi> i. max (b$$i) B)::'n::ordered_euclidean_space"
+ let ?a = "\<Sum>i\<in>Basis. min (a\<bullet>i) (-B) *\<^sub>R i::'n" and ?b = "\<Sum>i\<in>Basis. max (b\<bullet>i) B *\<^sub>R i::'n"
show "?f integrable_on {a..b}" apply(rule integrable_subinterval[of _ ?a ?b])
proof- have "ball 0 B \<subseteq> {?a..?b}" apply safe unfolding mem_ball mem_interval dist_norm
- proof case goal1 thus ?case using component_le_norm[of x i] by(auto simp add:field_simps) qed
+ proof case goal1 thus ?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]
thus "?f integrable_on {?a..?b}" unfolding integrable_on_def by auto
- show "{a..b} \<subseteq> {?a..?b}" apply safe unfolding mem_interval apply(rule,erule_tac x=i in allE) by auto qed
+ show "{a..b} \<subseteq> {?a..?b}" apply safe unfolding mem_interval apply(rule,erule_tac x=i in ballE) by auto qed
fix e::real assume "e>0" from as[OF this] guess B .. note B=conjunctD2[OF this,rule_format]
show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b} \<longrightarrow>
@@ -4584,14 +4517,15 @@
using B(2)[OF goal1(1)] B(2)[OF goal1(2)] by auto qed qed
next assume ?r note as = conjunctD2[OF this,rule_format]
- have "Cauchy (\<lambda>n. integral ({(\<chi>\<chi> i. - real n)::'n .. (\<chi>\<chi> i. real n)}) (\<lambda>x. if x \<in> s then f x else 0))"
+ let ?cube = "\<lambda>n. {(\<Sum>i\<in>Basis. - real n *\<^sub>R i)::'n .. (\<Sum>i\<in>Basis. real n *\<^sub>R i)} :: 'n set"
+ have "Cauchy (\<lambda>n. integral (?cube n) (\<lambda>x. if x \<in> s then f x else 0))"
proof(unfold Cauchy_def,safe) case goal1
from as(2)[OF this] guess B .. note B = conjunctD2[OF this,rule_format]
from real_arch_simple[of B] guess N .. note N = this
- { fix n assume n:"n \<ge> N" have "ball 0 B \<subseteq> {(\<chi>\<chi> i. - real n)::'n..\<chi>\<chi> i. real n}" apply safe
+ { fix n assume n:"n \<ge> N" have "ball 0 B \<subseteq> ?cube n" apply safe
unfolding mem_ball mem_interval dist_norm
- proof case goal1 thus ?case using component_le_norm[of x i]
- using n N by(auto simp add:field_simps) qed }
+ proof case goal1 thus ?case using Basis_le_norm[of i x] `i\<in>Basis`
+ using n N by(auto simp add:field_simps setsum_negf) qed }
thus ?case apply-apply(rule_tac x=N in exI) apply safe unfolding dist_norm apply(rule B(2)) by auto
qed from this[unfolded convergent_eq_cauchy[THEN sym]] guess i ..
note i = this[THEN LIMSEQ_D]
@@ -4611,9 +4545,9 @@
proof safe show "N \<le> n" using n by auto
fix x::"'n::ordered_euclidean_space" assume x:"x \<in> ball 0 B" hence "x\<in> ball 0 ?B" by auto
thus "x\<in>{a..b}" using ab by blast
- show "x\<in>{\<chi>\<chi> i. - real n..\<chi>\<chi> i. real n}" using x unfolding mem_interval mem_ball dist_norm apply-
- proof case goal1 thus ?case using component_le_norm[of x i]
- using n by(auto simp add:field_simps) qed qed qed qed qed
+ show "x\<in>?cube n" using x unfolding mem_interval mem_ball dist_norm apply-
+ proof case goal1 thus ?case using Basis_le_norm[of i x] `i\<in>Basis`
+ using n by(auto simp add:field_simps setsum_negf) qed qed qed qed qed
lemma integrable_altD: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on s"
@@ -4676,11 +4610,13 @@
note g = this(1) and this(2)[OF *] from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format]
note obt(2)[unfolded has_integral_alt'[of h]] note conjunctD2[OF this, rule_format]
note h = this(1) and this(2)[OF *] from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format]
- def c \<equiv> "(\<chi>\<chi> i. min (a$$i) (- (max B1 B2)))::'n" and d \<equiv> "(\<chi>\<chi> i. max (b$$i) (max B1 B2))::'n"
- have *:"ball 0 B1 \<subseteq> {c..d}" "ball 0 B2 \<subseteq> {c..d}" apply safe unfolding mem_ball mem_interval dist_norm
- proof(rule_tac[!] allI)
- case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def by auto next
- case goal2 thus ?case using component_le_norm[of x i] unfolding c_def d_def by auto qed
+ def c \<equiv> "\<Sum>i\<in>Basis. min (a\<bullet>i) (- (max B1 B2)) *\<^sub>R i::'n"
+ def d \<equiv> "\<Sum>i\<in>Basis. max (b\<bullet>i) (max B1 B2) *\<^sub>R i::'n"
+ have *:"ball 0 B1 \<subseteq> {c..d}" "ball 0 B2 \<subseteq> {c..d}"
+ apply safe unfolding mem_ball mem_interval dist_norm
+ proof(rule_tac[!] ballI)
+ case goal1 thus ?case using Basis_le_norm[of i x] unfolding c_def d_def by auto next
+ case goal2 thus ?case using Basis_le_norm[of i x] unfolding c_def d_def by auto qed
have **:"\<And>ch cg ag ah::real. norm(ah - ag) \<le> norm(ch - cg) \<Longrightarrow> norm(cg - i) < e / 4 \<Longrightarrow>
norm(ch - j) < e / 4 \<Longrightarrow> norm(ag - ah) < e"
using obt(3) unfolding real_norm_def by arith
@@ -4700,7 +4636,7 @@
unfolding integral_sub[OF h g,THEN sym] real_norm_def apply(subst **) defer apply(subst **) defer
apply(rule has_integral_subset_le) defer apply(rule integrable_integral integrable_sub h g)+
proof safe fix x assume "x\<in>{a..b}" thus "x\<in>{c..d}" unfolding mem_interval c_def d_def
- apply - apply rule apply(erule_tac x=i in allE) by auto
+ apply - apply rule apply(erule_tac x=i in ballE) by auto
qed(insert obt(4), auto) qed(insert obt(4), auto) qed note interv = this
show ?thesis unfolding integrable_alt[of f] apply safe apply(rule interv)
@@ -5019,7 +4955,7 @@
proof(case_tac[!] "content {a..b} = 0") assume as:"content {a..b} = 0"
show ?thesis using integrable_on_null[OF as] unfolding integral_null[OF as] using tendsto_const by auto
next assume ab:"content {a..b} \<noteq> 0"
- have fg:"\<forall>x\<in>{a..b}. \<forall> k. (f k x) $$ 0 \<le> (g x) $$ 0"
+ have fg:"\<forall>x\<in>{a..b}. \<forall> k. (f k x) \<bullet> 1 \<le> (g x) \<bullet> 1"
proof safe case goal1 note assms(3)[rule_format,OF this]
note * = Lim_component_ge[OF this trivial_limit_sequentially]
show ?case apply(rule *) unfolding eventually_sequentially
@@ -5030,10 +4966,10 @@
apply rule apply(rule integral_le) apply safe
apply(rule assms(1-2)[rule_format])+ using assms(4) by auto
then guess i .. note i=this
- have i':"\<And>k. (integral({a..b}) (f k)) \<le> i$$0"
+ have i':"\<And>k. (integral({a..b}) (f k)) \<le> i\<bullet>1"
apply(rule Lim_component_ge,rule i) apply(rule trivial_limit_sequentially)
unfolding eventually_sequentially apply(rule_tac x=k in exI)
- apply(rule transitive_stepwise_le) prefer 3 unfolding Eucl_real_simps apply(rule integral_le)
+ apply(rule transitive_stepwise_le) prefer 3 unfolding inner_simps real_inner_1_right apply(rule integral_le)
apply(rule assms(1-2)[rule_format])+ using assms(2) by auto
have "(g has_integral i) {a..b}" unfolding has_integral
@@ -5044,7 +4980,7 @@
apply(rule divide_pos_pos) by auto
from choice[OF this] guess c .. note c=conjunctD2[OF this[rule_format],rule_format]
- have "\<exists>r. \<forall>k\<ge>r. 0 \<le> i$$0 - (integral {a..b} (f k)) \<and> i$$0 - (integral {a..b} (f k)) < e / 4"
+ have "\<exists>r. \<forall>k\<ge>r. 0 \<le> i\<bullet>1 - (integral {a..b} (f k)) \<and> i\<bullet>1 - (integral {a..b} (f k)) < e / 4"
proof- case goal1 have "e/4 > 0" using e by auto
from LIMSEQ_D [OF i this] guess r ..
thus ?case apply(rule_tac x=r in exI) apply rule
@@ -5052,14 +4988,15 @@
proof- case goal1 thus ?case using i'[of k] by auto qed qed
then guess r .. note r=conjunctD2[OF this[rule_format]]
- have "\<forall>x\<in>{a..b}. \<exists>n\<ge>r. \<forall>k\<ge>n. 0 \<le> (g x)$$0 - (f k x)$$0 \<and>
- (g x)$$0 - (f k x)$$0 < e / (4 * content({a..b}))"
+ have "\<forall>x\<in>{a..b}. \<exists>n\<ge>r. \<forall>k\<ge>n. 0 \<le> (g x)\<bullet>1 - (f k x)\<bullet>1 \<and>
+ (g x)\<bullet>1 - (f k x)\<bullet>1 < e / (4 * content({a..b}))"
proof case goal1 have "e / (4 * content {a..b}) > 0" apply(rule divide_pos_pos,fact)
using ab content_pos_le[of a b] by auto
from assms(3)[rule_format, OF goal1, THEN LIMSEQ_D, OF this]
guess n .. note n=this
thus ?case apply(rule_tac x="n + r" in exI) apply safe apply(erule_tac[2-3] x=k in allE)
- unfolding dist_real_def using fg[rule_format,OF goal1] by(auto simp add:field_simps) qed
+ unfolding dist_real_def using fg[rule_format,OF goal1]
+ by (auto simp add:field_simps) qed
from bchoice[OF this] guess m .. note m=conjunctD2[OF this[rule_format],rule_format]
def d \<equiv> "\<lambda>x. c (m x) x"
@@ -5118,14 +5055,14 @@
next case goal3
note comb = integral_combine_tagged_division_topdown[OF assms(1)[rule_format] p(1)]
- have *:"\<And>sr sx ss ks kr::real. kr = sr \<longrightarrow> ks = ss \<longrightarrow> ks \<le> i \<and> sr \<le> sx \<and> sx \<le> ss \<and> 0 \<le> i$$0 - kr$$0
- \<and> i$$0 - kr$$0 < e/4 \<longrightarrow> abs(sx - i) < e/4" by auto
+ have *:"\<And>sr sx ss ks kr::real. kr = sr \<longrightarrow> ks = ss \<longrightarrow> ks \<le> i \<and> sr \<le> sx \<and> sx \<le> ss \<and> 0 \<le> i\<bullet>1 - kr\<bullet>1
+ \<and> i\<bullet>1 - kr\<bullet>1 < e/4 \<longrightarrow> abs(sx - i) < e/4" by auto
show ?case unfolding real_norm_def apply(rule *[rule_format],safe)
- apply(rule comb[of r],rule comb[of s]) apply(rule i'[unfolded Eucl_real_simps])
+ apply(rule comb[of r],rule comb[of s]) apply(rule i'[unfolded real_inner_1_right])
apply(rule_tac[1-2] setsum_mono) unfolding split_paired_all split_conv
apply(rule_tac[1-2] integral_le[OF ])
- proof safe show "0 \<le> i$$0 - (integral {a..b} (f r))$$0" using r(1) by auto
- show "i$$0 - (integral {a..b} (f r))$$0 < e / 4" using r(2) by auto
+ proof safe show "0 \<le> i\<bullet>1 - (integral {a..b} (f r))\<bullet>1" using r(1) by auto
+ show "i\<bullet>1 - (integral {a..b} (f r))\<bullet>1 < e / 4" using r(2) by auto
fix x k assume xk:"(x,k)\<in>p" from p'(4)[OF this] guess u v apply-by(erule exE)+ note uv=this
show "f r integrable_on k" "f s integrable_on k" "f (m x) integrable_on k" "f (m x) integrable_on k"
unfolding uv apply(rule_tac[!] integrable_on_subinterval[OF assms(1)[rule_format]])
@@ -5147,7 +5084,7 @@
\<forall>k. \<forall>x\<in>s. (f k x) \<le> (f (Suc k) x) \<Longrightarrow> \<forall>x\<in>s. ((\<lambda>k. f k x) ---> g x) sequentially \<Longrightarrow>
bounded {integral s (f k)| k. True} \<Longrightarrow> g integrable_on s \<and> ((\<lambda>k. integral s (f k)) ---> integral s g) sequentially"
proof- case goal1 note assms=this[rule_format]
- have "\<forall>x\<in>s. \<forall>k. (f k x)$$0 \<le> (g x)$$0" apply safe apply(rule Lim_component_ge)
+ have "\<forall>x\<in>s. \<forall>k. (f k x)\<bullet>1 \<le> (g x)\<bullet>1" apply safe apply(rule Lim_component_ge)
apply(rule goal1(4)[rule_format],assumption) apply(rule trivial_limit_sequentially)
unfolding eventually_sequentially apply(rule_tac x=k in exI)
apply(rule transitive_stepwise_le) using goal1(3) by auto note fg=this[rule_format]
@@ -5156,9 +5093,10 @@
apply(rule goal1(5)) apply(rule,rule integral_le) apply(rule goal1(2)[rule_format])+
using goal1(3) by auto then guess i .. note i=this
have "\<And>k. \<forall>x\<in>s. \<forall>n\<ge>k. f k x \<le> f n x" apply(rule,rule transitive_stepwise_le) using goal1(3) by auto
- hence i':"\<forall>k. (integral s (f k))$$0 \<le> i$$0" apply-apply(rule,rule Lim_component_ge)
+ hence i':"\<forall>k. (integral s (f k))\<bullet>1 \<le> i\<bullet>1" apply-apply(rule,rule Lim_component_ge)
apply(rule i,rule trivial_limit_sequentially) unfolding eventually_sequentially
apply(rule_tac x=k in exI,safe) apply(rule integral_component_le)
+ apply simp
apply(rule goal1(2)[rule_format])+ by auto
note int = assms(2)[unfolded integrable_alt[of _ s],THEN conjunct1,rule_format]
@@ -5199,14 +5137,14 @@
apply(rule norm_triangle_half_l) using B(2)[rule_format,OF ab] N[rule_format,of N]
apply-defer apply(subst norm_minus_commute) by auto
have *:"\<And>f1 f2 g. abs(f1 - i) < e / 2 \<longrightarrow> abs(f2 - g) < e / 2 \<longrightarrow> f1 \<le> f2 \<longrightarrow> f2 \<le> i
- \<longrightarrow> abs(g - i) < e" unfolding Eucl_real_simps by arith
+ \<longrightarrow> abs(g - i) < e" unfolding real_inner_1_right by arith
show "norm (integral {a..b} (\<lambda>x. if x \<in> s then g x else 0) - i) < e"
unfolding real_norm_def apply(rule *[rule_format])
apply(rule **[unfolded real_norm_def])
apply(rule M[rule_format,of "M + N",unfolded real_norm_def]) apply(rule le_add1)
apply(rule integral_le[OF int int]) defer
- apply(rule order_trans[OF _ i'[rule_format,of "M + N",unfolded Eucl_real_simps]])
- proof safe case goal2 have "\<And>m. x\<in>s \<Longrightarrow> \<forall>n\<ge>m. (f m x)$$0 \<le> (f n x)$$0"
+ apply(rule order_trans[OF _ i'[rule_format,of "M + N",unfolded real_inner_1_right]])
+ proof safe case goal2 have "\<And>m. x\<in>s \<Longrightarrow> \<forall>n\<ge>m. (f m x)\<bullet>1 \<le> (f n x)\<bullet>1"
apply(rule transitive_stepwise_le) using assms(3) by auto thus ?case by auto
next case goal1 show ?case apply(subst integral_restrict_univ[THEN sym,OF int])
unfolding ifif integral_restrict_univ[OF int']
@@ -5323,9 +5261,9 @@
lemma integral_norm_bound_integral_component: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
fixes g::"'n => 'b::ordered_euclidean_space"
- assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. norm(f x) \<le> (g x)$$k"
- shows "norm(integral s f) \<le> (integral s g)$$k"
-proof- have "norm (integral s f) \<le> integral s ((\<lambda>x. x $$ k) o g)"
+ assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. norm(f x) \<le> (g x)\<bullet>k"
+ shows "norm(integral s f) \<le> (integral s g)\<bullet>k"
+proof- have "norm (integral s f) \<le> integral s ((\<lambda>x. x \<bullet> k) o g)"
apply(rule integral_norm_bound_integral[OF assms(1)])
apply(rule integrable_linear[OF assms(2)],rule)
unfolding o_def by(rule assms)
@@ -5333,8 +5271,8 @@
lemma has_integral_norm_bound_integral_component: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
fixes g::"'n => 'b::ordered_euclidean_space"
- assumes "(f has_integral i) s" "(g has_integral j) s" "\<forall>x\<in>s. norm(f x) \<le> (g x)$$k"
- shows "norm(i) \<le> j$$k" using integral_norm_bound_integral_component[of f s g k]
+ assumes "(f has_integral i) s" "(g has_integral j) s" "\<forall>x\<in>s. norm(f x) \<le> (g x)\<bullet>k"
+ shows "norm(i) \<le> j\<bullet>k" using integral_norm_bound_integral_component[of f s g k]
unfolding integral_unique[OF assms(1)] integral_unique[OF assms(2)]
using assms by auto
@@ -5742,97 +5680,126 @@
shows "(\<lambda>x. setsum (\<lambda>a. f a x) t) absolutely_integrable_on s"
using assms(1,2) apply induct defer apply(subst setsum.insert) apply assumption+ by(rule,auto)
-lemma absolutely_integrable_vector_abs: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
- assumes "f absolutely_integrable_on s"
- shows "(\<lambda>x. (\<chi>\<chi> i. abs(f x$$i))::'c::ordered_euclidean_space) absolutely_integrable_on s"
-proof- have *:"\<And>x. ((\<chi>\<chi> i. abs(f x$$i))::'c::ordered_euclidean_space) = (setsum (\<lambda>i.
- (((\<lambda>y. (\<chi>\<chi> j. if j = i then y else 0)) o
- (((\<lambda>x. (norm((\<chi>\<chi> j. if j = i then x$$i else 0)::'c::ordered_euclidean_space))) o f))) x)) {..<DIM('c)})"
- unfolding euclidean_eq[where 'a='c] euclidean_component_setsum apply safe
- unfolding euclidean_lambda_beta'
- proof- case goal1 have *:"\<And>i xa. ((if i = xa then f x $$ xa else 0) * (if i = xa then f x $$ xa else 0)) =
- (if i = xa then (f x $$ xa) * (f x $$ xa) else 0)" by auto
- have *:"\<And>xa. norm ((\<chi>\<chi> j. if j = xa then f x $$ xa else 0)::'c) = (if xa<DIM('c) then abs (f x $$ xa) else 0)"
- unfolding norm_eq_sqrt_inner euclidean_inner[where 'a='c]
- by(auto simp add:setsum_delta[OF finite_lessThan] *)
- have "\<bar>f x $$ i\<bar> = (setsum (\<lambda>k. if k = i then abs ((f x)$$i) else 0) {..<DIM('c)})"
- unfolding setsum_delta[OF finite_lessThan] using goal1 by auto
- also have "... = (\<Sum>xa<DIM('c). ((\<lambda>y. (\<chi>\<chi> j. if j = xa then y else 0)::'c) \<circ>
- (\<lambda>x. (norm ((\<chi>\<chi> j. if j = xa then x $$ xa else 0)::'c))) \<circ> f) x $$ i)"
- unfolding o_def * apply(rule setsum_cong2)
- unfolding euclidean_lambda_beta'[OF goal1 ] by auto
- finally show ?case unfolding o_def . qed
- show ?thesis unfolding * apply(rule absolutely_integrable_setsum) apply(rule finite_lessThan)
- apply(rule absolutely_integrable_linear) unfolding o_def apply(rule absolutely_integrable_norm)
- apply(rule absolutely_integrable_linear[OF assms,unfolded o_def]) unfolding linear_linear
- apply(rule_tac[!] linearI) unfolding euclidean_eq[where 'a='c]
- by(auto simp:euclidean_component_scaleR[where 'a=real,unfolded real_scaleR_def])
+lemma bounded_linear_setsum:
+ fixes f :: "'i \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+ assumes f: "\<And>i. i\<in>I \<Longrightarrow> bounded_linear (f i)"
+ shows "bounded_linear (\<lambda>x. \<Sum>i\<in>I. f i x)"
+proof cases
+ assume "finite I" from this f show ?thesis
+ by (induct I) (auto intro!: bounded_linear_add bounded_linear_zero)
+qed (simp add: bounded_linear_zero)
+
+lemma absolutely_integrable_vector_abs:
+ fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
+ fixes T :: "'c::ordered_euclidean_space \<Rightarrow> 'b"
+ assumes f: "f absolutely_integrable_on s"
+ shows "(\<lambda>x. (\<Sum>i\<in>Basis. abs(f x\<bullet>T i) *\<^sub>R i)) absolutely_integrable_on s"
+ (is "?Tf absolutely_integrable_on s")
+proof -
+ have if_distrib: "\<And>P A B x. (if P then A else B) *\<^sub>R x = (if P then A *\<^sub>R x else B *\<^sub>R x)"
+ by simp
+ have *: "\<And>x. ?Tf x = (\<Sum>i\<in>Basis.
+ ((\<lambda>y. (\<Sum>j\<in>Basis. (if j = i then y else 0) *\<^sub>R j)) o
+ (\<lambda>x. (norm (\<Sum>j\<in>Basis. (if j = i then f x\<bullet>T i else 0) *\<^sub>R j)))) x)"
+ by (simp add: comp_def if_distrib setsum_cases)
+ show ?thesis
+ unfolding *
+ apply (rule absolutely_integrable_setsum[OF finite_Basis])
+ apply (rule absolutely_integrable_linear)
+ apply (rule absolutely_integrable_norm)
+ apply (rule absolutely_integrable_linear[OF f, unfolded o_def])
+ apply (auto simp: linear_linear euclidean_eq_iff[where 'a='c] inner_simps intro!: linearI)
+ done
qed
-lemma absolutely_integrable_max: fixes f g::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
+lemma absolutely_integrable_max:
+ fixes f g::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
- shows "(\<lambda>x. (\<chi>\<chi> i. max (f(x)$$i) (g(x)$$i))::'n::ordered_euclidean_space) absolutely_integrable_on s"
-proof- have *:"\<And>x. (1 / 2) *\<^sub>R (((\<chi>\<chi> i. \<bar>(f x - g x) $$ i\<bar>)::'n) + (f x + g x)) = (\<chi>\<chi> i. max (f(x)$$i) (g(x)$$i))"
- unfolding euclidean_eq[where 'a='n] by auto
+ shows "(\<lambda>x. (\<Sum>i\<in>Basis. max (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)::'n) absolutely_integrable_on s"
+proof -
+ have *:"\<And>x. (1 / 2) *\<^sub>R (((\<Sum>i\<in>Basis. \<bar>(f x - g x) \<bullet> i\<bar> *\<^sub>R i)::'n) + (f x + g x)) =
+ (\<Sum>i\<in>Basis. max (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)"
+ unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps)
note absolutely_integrable_sub[OF assms] absolutely_integrable_add[OF assms]
- note absolutely_integrable_vector_abs[OF this(1)] this(2)
- note absolutely_integrable_add[OF this] note absolutely_integrable_cmul[OF this,of "1/2"]
- thus ?thesis unfolding * . qed
-
-lemma absolutely_integrable_min: fixes f g::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
+ note absolutely_integrable_vector_abs[OF this(1), where T="\<lambda>x. x"] this(2)
+ note absolutely_integrable_add[OF this]
+ note absolutely_integrable_cmul[OF this, of "1/2"]
+ thus ?thesis unfolding * .
+qed
+
+lemma absolutely_integrable_min:
+ fixes f g::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
- shows "(\<lambda>x. (\<chi>\<chi> i. min (f(x)$$i) (g(x)$$i))::'n::ordered_euclidean_space) absolutely_integrable_on s"
-proof- have *:"\<And>x. (1 / 2) *\<^sub>R ((f x + g x) - ((\<chi>\<chi> i. \<bar>(f x - g x) $$ i\<bar>)::'n)) = (\<chi>\<chi> i. min (f(x)$$i) (g(x)$$i))"
- unfolding euclidean_eq[where 'a='n] by auto
+ shows "(\<lambda>x. (\<Sum>i\<in>Basis. min (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)::'n) absolutely_integrable_on s"
+proof -
+ have *:"\<And>x. (1 / 2) *\<^sub>R ((f x + g x) - (\<Sum>i\<in>Basis. \<bar>(f x - g x) \<bullet> i\<bar> *\<^sub>R i::'n)) =
+ (\<Sum>i\<in>Basis. min (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)"
+ unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps)
+
note absolutely_integrable_add[OF assms] absolutely_integrable_sub[OF assms]
- note this(1) absolutely_integrable_vector_abs[OF this(2)]
- note absolutely_integrable_sub[OF this] note absolutely_integrable_cmul[OF this,of "1/2"]
- thus ?thesis unfolding * . qed
-
-lemma absolutely_integrable_abs_eq: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
+ note this(1) absolutely_integrable_vector_abs[OF this(2), where T="\<lambda>x. x"]
+ note absolutely_integrable_sub[OF this]
+ note absolutely_integrable_cmul[OF this,of "1/2"]
+ thus ?thesis unfolding * .
+qed
+
+lemma absolutely_integrable_abs_eq:
+ fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
shows "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s \<and>
- (\<lambda>x. (\<chi>\<chi> i. abs(f x$$i))::'m) integrable_on s" (is "?l = ?r")
-proof assume ?l thus ?r apply-apply rule defer
+ (\<lambda>x. (\<Sum>i\<in>Basis. abs(f x\<bullet>i) *\<^sub>R i)::'m) integrable_on s" (is "?l = ?r")
+proof
+ assume ?l thus ?r apply-apply rule defer
apply(drule absolutely_integrable_vector_abs) by auto
-next assume ?r { presume lem:"\<And>f::'n \<Rightarrow> 'm. f integrable_on UNIV \<Longrightarrow>
- (\<lambda>x. (\<chi>\<chi> i. abs(f(x)$$i))::'m) integrable_on UNIV \<Longrightarrow> f absolutely_integrable_on UNIV"
- have *:"\<And>x. (\<chi>\<chi> i. \<bar>(if x \<in> s then f x else 0) $$ i\<bar>) = (if x\<in>s then (\<chi>\<chi> i. \<bar>f x $$ i\<bar>) else (0::'m))"
- unfolding euclidean_eq[where 'a='m] by auto
+next
+ assume ?r
+ { presume lem:"\<And>f::'n \<Rightarrow> 'm. f integrable_on UNIV \<Longrightarrow>
+ (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) integrable_on UNIV \<Longrightarrow> f absolutely_integrable_on UNIV"
+ have *:"\<And>x. (\<Sum>i\<in>Basis. \<bar>(if x \<in> s then f x else 0) \<bullet> i\<bar> *\<^sub>R i) =
+ (if x\<in>s then (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i) else (0::'m))"
+ unfolding euclidean_eq_iff[where 'a='m] by auto
show ?l apply(subst absolutely_integrable_restrict_univ[THEN sym]) apply(rule lem)
unfolding integrable_restrict_univ * using `?r` by auto }
- fix f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space" assume assms:"f integrable_on UNIV"
- "(\<lambda>x. (\<chi>\<chi> i. abs(f(x)$$i))::'m::ordered_euclidean_space) integrable_on UNIV"
- let ?B = "setsum (\<lambda>i. integral UNIV (\<lambda>x. (\<chi>\<chi> j. abs(f x$$j)) ::'m::ordered_euclidean_space) $$ i) {..<DIM('m)}"
+ fix f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
+ assume assms:"f integrable_on UNIV" "(\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) integrable_on UNIV"
+ let ?B = "\<Sum>i\<in>Basis. integral UNIV (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) \<bullet> i"
show "f absolutely_integrable_on UNIV"
apply(rule bounded_variation_absolutely_integrable[OF assms(1), where B="?B"],safe)
proof- case goal1 note d=this and d'=division_ofD[OF this]
have "(\<Sum>k\<in>d. norm (integral k f)) \<le>
- (\<Sum>k\<in>d. setsum (op $$ (integral k (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m))) {..<DIM('m)})" apply(rule setsum_mono)
+ (\<Sum>k\<in>d. setsum (op \<bullet> (integral k (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m))) Basis)"
+ apply(rule setsum_mono)
apply(rule order_trans[OF norm_le_l1]) apply(rule setsum_mono) unfolding lessThan_iff
- proof- fix k and i assume "k\<in>d" and i:"i<DIM('m)"
+ proof- fix k and i :: 'm assume "k\<in>d" and i:"i\<in>Basis"
from d'(4)[OF this(1)] guess a b apply-by(erule exE)+ note ab=this
- show "\<bar>integral k f $$ i\<bar> \<le> integral k (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m) $$ i" apply(rule abs_leI)
- unfolding euclidean_component_minus[THEN sym] defer apply(subst integral_neg[THEN sym])
- defer apply(rule_tac[1-2] integral_component_le) apply(rule integrable_neg)
+ show "\<bar>integral k f \<bullet> i\<bar> \<le> integral k (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) \<bullet> i"
+ apply (rule abs_leI)
+ unfolding inner_minus_left[THEN sym] defer apply(subst integral_neg[THEN sym])
+ 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] unfolding ab by auto
- qed also have "... \<le> setsum (op $$ (integral UNIV (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>))::'m)) {..<DIM('m)}"
+ integrable_on_subinterval[OF assms(2),of a b] i unfolding ab by auto
+ qed also have "... \<le> setsum (op \<bullet> (integral UNIV (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m))) Basis"
apply(subst setsum_commute,rule setsum_mono)
- proof- case goal1 have *:"(\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m) integrable_on \<Union>d"
+ proof- case goal1 have *:"(\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) integrable_on \<Union>d"
using integrable_on_subdivision[OF d assms(2)] by auto
- have "(\<Sum>i\<in>d. integral i (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m) $$ j)
- = integral (\<Union>d) (\<lambda>x. (\<chi>\<chi> j. abs(f x$$j)) ::'m::ordered_euclidean_space) $$ j"
- unfolding euclidean_component_setsum[THEN sym] integral_combine_division_topdown[OF * d] ..
- also have "... \<le> integral UNIV (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m) $$ j"
- apply(rule integral_subset_component_le) using assms * by auto
+ have "(\<Sum>i\<in>d. integral i (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j)
+ = integral (\<Union>d) (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j"
+ unfolding inner_setsum_left[symmetric] integral_combine_division_topdown[OF * d] ..
+ also have "... \<le> integral UNIV (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j"
+ apply(rule integral_subset_component_le) using assms * `j\<in>Basis` by auto
finally show ?case .
qed finally show ?case . qed qed
-lemma nonnegative_absolutely_integrable: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
- assumes "\<forall>x\<in>s. \<forall>i<DIM('m). 0 \<le> f(x)$$i" "f integrable_on s"
+lemma nonnegative_absolutely_integrable:
+ fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
+ assumes "\<forall>x\<in>s. \<forall>i\<in>Basis. 0 \<le> f(x)\<bullet>i" "f integrable_on s"
shows "f absolutely_integrable_on s"
- unfolding absolutely_integrable_abs_eq apply rule defer
- apply(rule integrable_eq[of _ f]) using assms apply-apply(subst euclidean_eq) by auto
+ unfolding absolutely_integrable_abs_eq
+ apply rule
+ apply (rule assms)
+ apply (rule integrable_eq[of _ f])
+ using assms
+ apply (auto simp: euclidean_eq_iff[where 'a='m])
+ done
lemma absolutely_integrable_integrable_bound: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
assumes "\<forall>x\<in>s. norm(f x) \<le> g x" "f integrable_on s" "g integrable_on s"
@@ -5881,8 +5848,8 @@
apply(subst Inf_insert_finite) apply(rule finite_imageI[OF insert(1)])
proof(cases "k={}") case True
thus ?P apply(subst if_P) defer apply(rule insert(5)[rule_format]) by auto
- next case False thus ?P apply(subst if_not_P) defer
- apply(rule absolutely_integrable_min[where 'n=real,unfolded Eucl_real_simps])
+ next case False thus ?P apply(subst if_not_P) defer
+ apply (rule absolutely_integrable_min[where 'n=real, unfolded Basis_real_def, simplified])
defer apply(rule insert(3)[OF False]) using insert(5) by auto
qed qed auto
@@ -5897,7 +5864,7 @@
proof(cases "k={}") case True
thus ?P apply(subst if_P) defer apply(rule insert(5)[rule_format]) by auto
next case False thus ?P apply(subst if_not_P) defer
- apply(rule absolutely_integrable_max[where 'n=real,unfolded Eucl_real_simps])
+ apply (rule absolutely_integrable_max[where 'n=real, unfolded Basis_real_def, simplified])
defer apply(rule insert(3)[OF False]) using insert(5) by auto
qed qed auto
--- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Dec 14 15:46:01 2012 +0100
@@ -398,6 +398,53 @@
then show "h = g" by (simp add: ext)
qed
+text {* TODO: The following lemmas about adjoints should hold for any
+Hilbert space (i.e. complete inner product space).
+(see \url{http://en.wikipedia.org/wiki/Hermitian_adjoint})
+*}
+
+lemma adjoint_works:
+ fixes f:: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+ assumes lf: "linear f"
+ shows "x \<bullet> adjoint f y = f x \<bullet> y"
+proof -
+ have "\<forall>y. \<exists>w. \<forall>x. f x \<bullet> y = x \<bullet> w"
+ proof (intro allI exI)
+ fix y :: "'m" and x
+ let ?w = "(\<Sum>i\<in>Basis. (f i \<bullet> y) *\<^sub>R i) :: 'n"
+ have "f x \<bullet> y = f (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i) \<bullet> y"
+ by (simp add: euclidean_representation)
+ also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R f i) \<bullet> y"
+ unfolding linear_setsum[OF lf finite_Basis]
+ by (simp add: linear_cmul[OF lf])
+ finally show "f x \<bullet> y = x \<bullet> ?w"
+ by (simp add: inner_setsum_left inner_setsum_right mult_commute)
+ qed
+ then show ?thesis
+ unfolding adjoint_def choice_iff
+ by (intro someI2_ex[where Q="\<lambda>f'. x \<bullet> f' y = f x \<bullet> y"]) auto
+qed
+
+lemma adjoint_clauses:
+ fixes f:: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+ assumes lf: "linear f"
+ shows "x \<bullet> adjoint f y = f x \<bullet> y"
+ and "adjoint f y \<bullet> x = y \<bullet> f x"
+ by (simp_all add: adjoint_works[OF lf] inner_commute)
+
+lemma adjoint_linear:
+ fixes f:: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+ assumes lf: "linear f"
+ shows "linear (adjoint f)"
+ by (simp add: lf linear_def euclidean_eq_iff[where 'a='n] euclidean_eq_iff[where 'a='m]
+ adjoint_clauses[OF lf] inner_simps)
+
+lemma adjoint_adjoint:
+ fixes f:: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+ assumes lf: "linear f"
+ shows "adjoint (adjoint f) = f"
+ by (rule adjoint_unique, simp add: adjoint_clauses [OF lf])
+
subsection {* Interlude: Some properties of real sets *}
lemma seq_mono_lemma: assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n" and "\<forall>n \<ge> m. e n <= e m"
@@ -1261,77 +1308,37 @@
subsection{* Euclidean Spaces as Typeclass*}
-lemma independent_eq_inj_on:
- fixes D :: nat
- and f :: "nat \<Rightarrow> 'c::real_vector"
- assumes "inj_on f {..<D}"
- shows "independent (f ` {..<D}) \<longleftrightarrow> (\<forall>a u. a < D \<longrightarrow> (\<Sum>i\<in>{..<D}-{a}. u (f i) *\<^sub>R f i) \<noteq> f a)"
-proof -
- from assms have eq: "\<And>i. i < D \<Longrightarrow> f ` {..<D} - {f i} = f`({..<D} - {i})"
- and inj: "\<And>i. inj_on f ({..<D} - {i})"
- by (auto simp: inj_on_def)
- have *: "\<And>i. finite (f ` {..<D} - {i})" by simp
- show ?thesis unfolding dependent_def span_finite[OF *]
- by (auto simp: eq setsum_reindex[OF inj])
-qed
-
-lemma independent_basis: "independent (basis ` {..<DIM('a)} :: 'a::euclidean_space set)"
- unfolding independent_eq_inj_on [OF basis_inj]
+lemma independent_Basis: "independent Basis"
+ unfolding dependent_def
+ apply (subst span_finite)
+ apply simp
apply clarify
- apply (drule_tac f="inner (basis a)" in arg_cong)
- apply (simp add: inner_setsum_right dot_basis)
+ apply (drule_tac f="inner a" in arg_cong)
+ apply (simp add: inner_Basis inner_setsum_right eq_commute)
+ done
+
+lemma span_Basis[simp]: "span Basis = (UNIV :: 'a::euclidean_space set)"
+ apply (subst span_finite)
+ apply simp
+ apply (safe intro!: UNIV_I)
+ apply (rule_tac x="inner x" in exI)
+ apply (simp add: euclidean_representation)
done
-lemma (in euclidean_space) range_basis: "range basis = insert 0 (basis ` {..<DIM('a)})"
-proof -
- have *: "UNIV = {..<DIM('a)} \<union> {DIM('a)..}" by auto
- show ?thesis unfolding * image_Un basis_finite by auto
-qed
-
-lemma (in euclidean_space) range_basis_finite[intro]: "finite (range basis)"
- unfolding range_basis by auto
-
-lemma span_basis: "span (range basis) = (UNIV :: 'a::euclidean_space set)"
-proof -
- { fix x :: 'a
- have "(\<Sum>i<DIM('a). (x $$ i) *\<^sub>R basis i) \<in> span (range basis :: 'a set)"
- by (simp add: span_setsum span_mul span_superset)
- then have "x \<in> span (range basis)"
- by (simp only: euclidean_representation [symmetric])
- } then show ?thesis by auto
-qed
-
-lemma basis_representation:
- "\<exists>u. x = (\<Sum>v\<in>basis ` {..<DIM('a)}. u v *\<^sub>R (v\<Colon>'a\<Colon>euclidean_space))"
-proof -
- have "x\<in>UNIV" by auto from this[unfolded span_basis[THEN sym]]
- have "\<exists>u. (\<Sum>v\<in>basis ` {..<DIM('a)}. u v *\<^sub>R v) = x"
- unfolding range_basis span_insert_0 apply(subst (asm) span_finite) by auto
- then show ?thesis by fastforce
-qed
-
-lemma span_basis'[simp]:"span ((basis::nat=>'a) ` {..<DIM('a::euclidean_space)}) = UNIV"
- apply(subst span_basis[symmetric])
- unfolding range_basis
- apply auto
- done
-
-lemma card_basis[simp]:"card ((basis::nat=>'a) ` {..<DIM('a::euclidean_space)}) = DIM('a)"
- apply (subst card_image)
- using basis_inj apply auto
- done
-
-lemma in_span_basis: "(x::'a::euclidean_space) \<in> span (basis ` {..<DIM('a)})"
- unfolding span_basis' ..
-
-lemma norm_bound_component_le: "norm (x::'a::euclidean_space) \<le> e \<Longrightarrow> \<bar>x$$i\<bar> <= e"
- by (metis component_le_norm order_trans)
-
-lemma norm_bound_component_lt: "norm (x::'a::euclidean_space) < e \<Longrightarrow> \<bar>x$$i\<bar> < e"
- by (metis component_le_norm basic_trans_rules(21))
-
-lemma norm_le_l1: "norm (x::'a::euclidean_space) \<le> (\<Sum>i<DIM('a). \<bar>x $$ i\<bar>)"
- apply (subst euclidean_representation[of x])
+lemma in_span_Basis: "x \<in> span Basis"
+ unfolding span_Basis ..
+
+lemma Basis_le_norm: "b \<in> Basis \<Longrightarrow> \<bar>x \<bullet> b\<bar> \<le> norm x"
+ by (rule order_trans [OF Cauchy_Schwarz_ineq2]) simp
+
+lemma norm_bound_Basis_le: "b \<in> Basis \<Longrightarrow> norm x \<le> e \<Longrightarrow> \<bar>x \<bullet> b\<bar> \<le> e"
+ by (metis Basis_le_norm order_trans)
+
+lemma norm_bound_Basis_lt: "b \<in> Basis \<Longrightarrow> norm x < e \<Longrightarrow> \<bar>x \<bullet> b\<bar> < e"
+ by (metis Basis_le_norm basic_trans_rules(21))
+
+lemma norm_le_l1: "norm x \<le> (\<Sum>b\<in>Basis. \<bar>x \<bullet> b\<bar>)"
+ apply (subst euclidean_representation[of x, symmetric])
apply (rule order_trans[OF norm_setsum])
apply (auto intro!: setsum_mono)
done
@@ -1339,61 +1346,29 @@
lemma setsum_norm_allsubsets_bound:
fixes f:: "'a \<Rightarrow> 'n::euclidean_space"
assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"
- shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real DIM('n) * e"
+ shows "(\<Sum>x\<in>P. norm (f x)) \<le> 2 * real DIM('n) * e"
proof -
- let ?d = "real DIM('n)"
- let ?nf = "\<lambda>x. norm (f x)"
- let ?U = "{..<DIM('n)}"
- have th0: "setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $$ i\<bar>) ?U) P = setsum (\<lambda>i. setsum (\<lambda>x. \<bar>f x $$ i\<bar>) P) ?U"
+ have "(\<Sum>x\<in>P. norm (f x)) \<le> (\<Sum>x\<in>P. \<Sum>b\<in>Basis. \<bar>f x \<bullet> b\<bar>)"
+ by (rule setsum_mono) (rule norm_le_l1)
+ also have "(\<Sum>x\<in>P. \<Sum>b\<in>Basis. \<bar>f x \<bullet> b\<bar>) = (\<Sum>b\<in>Basis. \<Sum>x\<in>P. \<bar>f x \<bullet> b\<bar>)"
by (rule setsum_commute)
- have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def)
- have "setsum ?nf P \<le> setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $$ i\<bar>) ?U) P"
- by (rule setsum_mono) (rule norm_le_l1)
- also have "\<dots> \<le> 2 * ?d * e"
- unfolding th0 th1
+ also have "\<dots> \<le> of_nat (card (Basis :: 'n set)) * (2 * e)"
proof (rule setsum_bounded)
- fix i assume i: "i \<in> ?U"
- let ?Pp = "{x. x\<in> P \<and> f x $$ i \<ge> 0}"
- let ?Pn = "{x. x \<in> P \<and> f x $$ i < 0}"
- have thp: "P = ?Pp \<union> ?Pn" by auto
- have thp0: "?Pp \<inter> ?Pn ={}" by auto
- have PpP: "?Pp \<subseteq> P" and PnP: "?Pn \<subseteq> P" by blast+
- have Ppe:"setsum (\<lambda>x. \<bar>f x $$ i\<bar>) ?Pp \<le> e"
- using component_le_norm[of "setsum (\<lambda>x. f x) ?Pp" i] fPs[OF PpP]
- by(auto intro: abs_le_D1)
- have Pne: "setsum (\<lambda>x. \<bar>f x $$ i\<bar>) ?Pn \<le> e"
- using component_le_norm[of "setsum (\<lambda>x. - f x) ?Pn" i] fPs[OF PnP]
- by(auto simp add: setsum_negf intro: abs_le_D1)
- have "setsum (\<lambda>x. \<bar>f x $$ i\<bar>) P = setsum (\<lambda>x. \<bar>f x $$ i\<bar>) ?Pp + setsum (\<lambda>x. \<bar>f x $$ i\<bar>) ?Pn"
- apply (subst thp)
- apply (rule setsum_Un_zero)
- using fP thp0 apply auto
- done
- also have "\<dots> \<le> 2*e" using Pne Ppe by arith
- finally show "setsum (\<lambda>x. \<bar>f x $$ i\<bar>) P \<le> 2*e" .
+ fix i :: 'n assume i: "i \<in> Basis"
+ have "norm (\<Sum>x\<in>P. \<bar>f x \<bullet> i\<bar>) \<le>
+ norm ((\<Sum>x\<in>P \<inter> - {x. f x \<bullet> i < 0}. f x) \<bullet> i) + norm ((\<Sum>x\<in>P \<inter> {x. f x \<bullet> i < 0}. f x) \<bullet> i)"
+ by (simp add: abs_real_def setsum_cases[OF fP] setsum_negf uminus_add_conv_diff
+ norm_triangle_ineq4 inner_setsum_left
+ del: real_norm_def)
+ also have "\<dots> \<le> e + e" unfolding real_norm_def
+ by (intro add_mono norm_bound_Basis_le i fPs) auto
+ finally show "(\<Sum>x\<in>P. \<bar>f x \<bullet> i\<bar>) \<le> 2*e" by simp
qed
+ also have "\<dots> = 2 * real DIM('n) * e"
+ by (simp add: real_of_nat_def)
finally show ?thesis .
qed
-lemma lambda_skolem': "(\<forall>i<DIM('a::euclidean_space). \<exists>x. P i x) \<longleftrightarrow>
- (\<exists>x::'a. \<forall>i<DIM('a). P i (x$$i))" (is "?lhs \<longleftrightarrow> ?rhs")
-proof -
- let ?S = "{..<DIM('a)}"
- { assume H: "?rhs"
- then have ?lhs by auto }
- moreover
- { assume H: "?lhs"
- then obtain f where f:"\<forall>i<DIM('a). P i (f i)" unfolding choice_iff' by metis
- let ?x = "(\<chi>\<chi> i. (f i)) :: 'a"
- { fix i assume i:"i<DIM('a)"
- with f have "P i (f i)" by metis
- then have "P i (?x$$i)" using i by auto }
- then have "\<forall>i<DIM('a). P i (?x$$i)" by metis
- then have ?rhs by metis }
- ultimately show ?thesis by metis
-qed
-
-
subsection {* Linearity and Bilinearity continued *}
lemma linear_bounded:
@@ -1401,29 +1376,25 @@
assumes lf: "linear f"
shows "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"
proof -
- let ?S = "{..<DIM('a)}"
- let ?B = "setsum (\<lambda>i. norm(f(basis i))) ?S"
- have fS: "finite ?S" by simp
+ let ?B = "\<Sum>b\<in>Basis. norm (f b)"
{ fix x:: "'a"
- let ?g = "(\<lambda> i. (x$$i) *\<^sub>R (basis i) :: 'a)"
- have "norm (f x) = norm (f (setsum (\<lambda>i. (x$$i) *\<^sub>R (basis i)) ?S))"
- apply (subst euclidean_representation[of x])
- apply rule
- done
- also have "\<dots> = norm (setsum (\<lambda> i. (x$$i) *\<^sub>R f (basis i)) ?S)"
- using linear_setsum[OF lf fS, of ?g, unfolded o_def] linear_cmul[OF lf] by auto
- finally have th0: "norm (f x) = norm (setsum (\<lambda>i. (x$$i) *\<^sub>R f (basis i))?S)" .
- { fix i assume i: "i \<in> ?S"
- from component_le_norm[of x i]
- have "norm ((x$$i) *\<^sub>R f (basis i :: 'a)) \<le> norm (f (basis i)) * norm x"
+ let ?g = "\<lambda>b. (x \<bullet> b) *\<^sub>R f b"
+ have "norm (f x) = norm (f (\<Sum>b\<in>Basis. (x \<bullet> b) *\<^sub>R b))"
+ unfolding euclidean_representation ..
+ also have "\<dots> = norm (setsum ?g Basis)"
+ using linear_setsum[OF lf finite_Basis, of "\<lambda>b. (x \<bullet> b) *\<^sub>R b", unfolded o_def] linear_cmul[OF lf] by auto
+ finally have th0: "norm (f x) = norm (setsum ?g Basis)" .
+ { fix i :: 'a assume i: "i \<in> Basis"
+ from Basis_le_norm[OF i, of x]
+ have "norm (?g i) \<le> norm (f i) * norm x"
unfolding norm_scaleR
- apply (simp only: mult_commute)
+ apply (subst mult_commute)
apply (rule mult_mono)
apply (auto simp add: field_simps)
done }
- then have th: "\<forall>i\<in> ?S. norm ((x$$i) *\<^sub>R f (basis i :: 'a)) \<le> norm (f (basis i)) * norm x"
+ then have th: "\<forall>b\<in>Basis. norm (?g b) \<le> norm (f b) * norm x"
by metis
- from setsum_norm_le[of _ "\<lambda>i. (x$$i) *\<^sub>R (f (basis i))", OF th]
+ from setsum_norm_le[of _ ?g, OF th]
have "norm (f x) \<le> ?B * norm x" unfolding th0 setsum_left_distrib by metis}
then show ?thesis by blast
qed
@@ -1438,12 +1409,15 @@
let ?K = "\<bar>B\<bar> + 1"
have Kp: "?K > 0" by arith
{ assume C: "B < 0"
- have "((\<chi>\<chi> i. 1)::'a) \<noteq> 0" unfolding euclidean_eq[where 'a='a]
- by(auto intro!:exI[where x=0])
- then have "norm ((\<chi>\<chi> i. 1)::'a) > 0" by auto
- with C have "B * norm ((\<chi>\<chi> i. 1)::'a) < 0"
+ def One \<equiv> "\<Sum>Basis ::'a"
+ then have "One \<noteq> 0"
+ unfolding euclidean_eq_iff[where 'a='a]
+ by (simp add: inner_setsum_left inner_Basis setsum_cases)
+ then have "norm One > 0" by auto
+ with C have "B * norm One < 0"
by (simp add: mult_less_0_iff)
- with B[rule_format, of "(\<chi>\<chi> i. 1)::'a"] norm_ge_zero[of "f ((\<chi>\<chi> i. 1)::'a)"] have False by simp
+ with B[rule_format, of One] norm_ge_zero[of "f One"]
+ have False by simp
}
then have Bp: "B \<ge> 0" by (metis not_leE)
{ fix x::"'a"
@@ -1492,33 +1466,27 @@
fixes h:: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'k::real_normed_vector"
assumes bh: "bilinear h"
shows "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
-proof -
- let ?M = "{..<DIM('m)}"
- let ?N = "{..<DIM('n)}"
- let ?B = "setsum (\<lambda>(i,j). norm (h (basis i) (basis j))) (?M \<times> ?N)"
- have fM: "finite ?M" and fN: "finite ?N" by simp_all
- { fix x:: "'m" and y :: "'n"
- have "norm (h x y) = norm (h (setsum (\<lambda>i. (x$$i) *\<^sub>R basis i) ?M) (setsum (\<lambda>i. (y$$i) *\<^sub>R basis i) ?N))"
- apply(subst euclidean_representation[where 'a='m])
- apply(subst euclidean_representation[where 'a='n])
- apply rule
- done
- also have "\<dots> = norm (setsum (\<lambda> (i,j). h ((x$$i) *\<^sub>R basis i) ((y$$j) *\<^sub>R basis j)) (?M \<times> ?N))"
- unfolding bilinear_setsum[OF bh fM fN] ..
- finally have th: "norm (h x y) = \<dots>" .
- have "norm (h x y) \<le> ?B * norm x * norm y"
- apply (simp add: setsum_left_distrib th)
+proof (clarify intro!: exI[of _ "\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)"])
+ fix x:: "'m" and y :: "'n"
+ have "norm (h x y) = norm (h (setsum (\<lambda>i. (x \<bullet> i) *\<^sub>R i) Basis) (setsum (\<lambda>i. (y \<bullet> i) *\<^sub>R i) Basis))"
+ apply(subst euclidean_representation[where 'a='m])
+ apply(subst euclidean_representation[where 'a='n])
+ apply rule
+ done
+ also have "\<dots> = norm (setsum (\<lambda> (i,j). h ((x \<bullet> i) *\<^sub>R i) ((y \<bullet> j) *\<^sub>R j)) (Basis \<times> Basis))"
+ unfolding bilinear_setsum[OF bh finite_Basis finite_Basis] ..
+ finally have th: "norm (h x y) = \<dots>" .
+ show "norm (h x y) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)) * norm x * norm y"
+ apply (auto simp add: setsum_left_distrib th setsum_cartesian_product)
apply (rule setsum_norm_le)
- using fN fM
apply simp
apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh]
field_simps simp del: scaleR_scaleR)
apply (rule mult_mono)
- apply (auto simp add: zero_le_mult_iff component_le_norm)
+ apply (auto simp add: zero_le_mult_iff Basis_le_norm)
apply (rule mult_mono)
- apply (auto simp add: zero_le_mult_iff component_le_norm)
- done }
- then show ?thesis by metis
+ apply (auto simp add: zero_le_mult_iff Basis_le_norm)
+ done
qed
lemma bilinear_bounded_pos:
@@ -1582,8 +1550,8 @@
lemma independent_bound:
fixes S:: "('a::euclidean_space) set"
- shows "independent S \<Longrightarrow> finite S \<and> card S <= DIM('a::euclidean_space)"
- using independent_span_bound[of "(basis::nat=>'a) ` {..<DIM('a)}" S] by auto
+ shows "independent S \<Longrightarrow> finite S \<and> card S \<le> DIM('a::euclidean_space)"
+ using independent_span_bound[OF finite_Basis, of S] by auto
lemma dependent_biggerset:
"(finite (S::('a::euclidean_space) set) ==> card S > DIM('a)) ==> dependent S"
@@ -1666,9 +1634,8 @@
text {* More lemmas about dimension. *}
lemma dim_UNIV: "dim (UNIV :: ('a::euclidean_space) set) = DIM('a)"
- apply (rule dim_unique[of "(basis::nat=>'a) ` {..<DIM('a)}"])
- using independent_basis apply auto
- done
+ using independent_Basis
+ by (intro dim_unique[of Basis]) auto
lemma dim_subset:
"(S:: ('a::euclidean_space) set) \<subseteq> T \<Longrightarrow> dim S \<le> dim T"
@@ -2256,20 +2223,9 @@
lemma linear_eq_stdbasis:
assumes lf: "linear (f::'a::euclidean_space \<Rightarrow> _)"
and lg: "linear g"
- and fg: "\<forall>i<DIM('a::euclidean_space). f (basis i) = g(basis i)"
+ and fg: "\<forall>b\<in>Basis. f b = g b"
shows "f = g"
-proof -
- let ?U = "{..<DIM('a)}"
- let ?I = "(basis::nat=>'a) ` {..<DIM('a)}"
- { fix x assume x: "x \<in> (UNIV :: 'a set)"
- from equalityD2[OF span_basis'[where 'a='a]]
- have IU: " (UNIV :: 'a set) \<subseteq> span ?I" by blast
- have "f x = g x"
- apply (rule linear_eq[OF lf lg IU,rule_format])
- using fg x apply auto
- done
- } then show ?thesis by auto
-qed
+ using linear_eq[OF lf lg, of _ Basis] fg by auto
text {* Similar results for bilinear functions. *}
@@ -2303,14 +2259,9 @@
fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> _"
assumes bf: "bilinear f"
and bg: "bilinear g"
- and fg: "\<forall>i<DIM('a). \<forall>j<DIM('b). f (basis i) (basis j) = g (basis i) (basis j)"
+ and fg: "\<forall>i\<in>Basis. \<forall>j\<in>Basis. f i j = g i j"
shows "f = g"
-proof -
- from fg have th: "\<forall>x \<in> (basis ` {..<DIM('a)}). \<forall>y\<in> (basis ` {..<DIM('b)}). f x y = g x y"
- by blast
- from bilinear_eq[OF bf bg equalityD2[OF span_basis'] equalityD2[OF span_basis'] th]
- show ?thesis by blast
-qed
+ using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis] fg] by blast
text {* Detailed theorems about left and right invertibility in general case. *}
@@ -2319,10 +2270,10 @@
assumes lf: "linear f" and fi: "inj f"
shows "\<exists>g. linear g \<and> g o f = id"
proof -
- from linear_independent_extend[OF independent_injective_image, OF independent_basis, OF lf fi]
+ from linear_independent_extend[OF independent_injective_image, OF independent_Basis, OF lf fi]
obtain h:: "'b => 'a" where
- h: "linear h" "\<forall>x \<in> f ` basis ` {..<DIM('a)}. h x = inv f x" by blast
- from h(2) have th: "\<forall>i<DIM('a). (h \<circ> f) (basis i) = id (basis i)"
+ h: "linear h" "\<forall>x \<in> f ` Basis. h x = inv f x" by blast
+ from h(2) have th: "\<forall>i\<in>Basis. (h \<circ> f) i = id i"
using inv_o_cancel[OF fi, unfolded fun_eq_iff id_def o_def]
by auto
@@ -2336,12 +2287,12 @@
assumes lf: "linear f" and sf: "surj f"
shows "\<exists>g. linear g \<and> f o g = id"
proof -
- from linear_independent_extend[OF independent_basis[where 'a='b],of "inv f"]
+ from linear_independent_extend[OF independent_Basis[where 'a='b],of "inv f"]
obtain h:: "'b \<Rightarrow> 'a" where
- h: "linear h" "\<forall> x\<in> basis ` {..<DIM('b)}. h x = inv f x" by blast
+ h: "linear h" "\<forall>x\<in>Basis. h x = inv f x" by blast
from h(2)
- have th: "\<forall>i<DIM('b). (f o h) (basis i) = id (basis i)"
- using sf by(auto simp add: surj_iff_all)
+ have th: "\<forall>i\<in>Basis. (f o h) i = id i"
+ using sf by (auto simp add: surj_iff_all)
from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th]
have "f o h = id" .
then show ?thesis using h(1) by blast
@@ -2538,18 +2489,18 @@
subsection {* Infinity norm *}
-definition "infnorm (x::'a::euclidean_space) = Sup {abs(x$$i) |i. i<DIM('a)}"
+definition "infnorm (x::'a::euclidean_space) = Sup { abs (x \<bullet> b) |b. b \<in> Basis}"
lemma numseg_dimindex_nonempty: "\<exists>i. i \<in> (UNIV :: 'n set)"
by auto
lemma infnorm_set_image:
- "{abs((x::'a::euclidean_space)$$i) |i. i<DIM('a)} =
- (\<lambda>i. abs(x$$i)) ` {..<DIM('a)}" by blast
+ "{ abs ((x::'a::euclidean_space) \<bullet> i) |i. i \<in> Basis} = (\<lambda>i. abs(x \<bullet> i)) ` Basis"
+ by blast
lemma infnorm_set_lemma:
- shows "finite {abs((x::'a::euclidean_space)$$i) |i. i<DIM('a)}"
- and "{abs(x$$i) |i. i<DIM('a::euclidean_space)} \<noteq> {}"
+ shows "finite {abs((x::'a::euclidean_space) \<bullet> i) |i. i \<in> Basis}"
+ and "{abs(x \<bullet> i) |i. i \<in> Basis} \<noteq> {}"
unfolding infnorm_set_image
by auto
@@ -2557,25 +2508,26 @@
unfolding infnorm_def
unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]
unfolding infnorm_set_image
- by auto
+ by (auto simp: ex_in_conv)
lemma infnorm_triangle: "infnorm ((x::'a::euclidean_space) + y) \<le> infnorm x + infnorm y"
proof -
have th: "\<And>x y (z::real). x - y <= z \<longleftrightarrow> x - z <= y" by arith
have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
have th2: "\<And>x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith
- have *:"\<And>i. i \<in> {..<DIM('a)} \<longleftrightarrow> i <DIM('a)" by auto
show ?thesis
- unfolding infnorm_def unfolding Sup_finite_le_iff[ OF infnorm_set_lemma[where 'a='a]]
- apply (subst diff_le_eq[symmetric])
- unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]
- unfolding infnorm_set_image bex_simps
- apply (subst th)
- unfolding th1 *
- unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma[where 'a='a]]
- unfolding infnorm_set_image ball_simps bex_simps
- unfolding euclidean_simps apply (metis th2)
- done
+ unfolding infnorm_def
+ unfolding Sup_finite_le_iff[ OF infnorm_set_lemma[where 'a='a]]
+ apply (subst diff_le_eq[symmetric])
+ unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]
+ unfolding infnorm_set_image bex_simps
+ apply (subst th)
+ unfolding th1
+ unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma[where 'a='a]]
+ unfolding infnorm_set_image ball_simps bex_simps
+ apply (simp add: inner_add_left)
+ apply (metis th2)
+ done
qed
lemma infnorm_eq_0: "infnorm x = 0 \<longleftrightarrow> (x::_::euclidean_space) = 0"
@@ -2584,7 +2536,7 @@
unfolding infnorm_def
unfolding Sup_finite_le_iff[OF infnorm_set_lemma]
unfolding infnorm_set_image ball_simps
- apply (subst (1) euclidean_eq)
+ apply (subst (1) euclidean_eq_iff)
apply auto
done
then show ?thesis using infnorm_pos_le[of x] by simp
@@ -2620,29 +2572,22 @@
lemma real_abs_infnorm: " \<bar>infnorm x\<bar> = infnorm x"
using infnorm_pos_le[of x] by arith
-lemma component_le_infnorm: "\<bar>x$$i\<bar> \<le> infnorm (x::'a::euclidean_space)"
-proof (cases "i<DIM('a)")
- case False
- then show ?thesis using infnorm_pos_le by auto
-next
- case True
- let ?U = "{..<DIM('a)}"
- let ?S = "{\<bar>x$$i\<bar> |i. i<DIM('a)}"
- have fS: "finite ?S" unfolding image_Collect[symmetric]
- apply (rule finite_imageI) apply simp done
- have S0: "?S \<noteq> {}" by blast
- have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
- show ?thesis unfolding infnorm_def
- apply(subst Sup_finite_ge_iff) using Sup_finite_in[OF fS S0]
- using infnorm_set_image using True apply auto
- done
+lemma Basis_le_infnorm:
+ assumes b: "b \<in> Basis" shows "\<bar>x \<bullet> b\<bar> \<le> infnorm (x::'a::euclidean_space)"
+ unfolding infnorm_def
+proof (subst Sup_finite_ge_iff)
+ let ?S = "{\<bar>x \<bullet> i\<bar> |i. i \<in> Basis}"
+ show "finite ?S" by (rule infnorm_set_lemma)
+ show "?S \<noteq> {}" by auto
+ show "Bex ?S (op \<le> \<bar>x \<bullet> b\<bar>)"
+ using b by (auto intro!: exI[of _ b])
qed
lemma infnorm_mul_lemma: "infnorm(a *\<^sub>R x) <= \<bar>a\<bar> * infnorm x"
apply (subst infnorm_def)
unfolding Sup_finite_le_iff[OF infnorm_set_lemma]
- unfolding infnorm_set_image ball_simps euclidean_component_scaleR abs_mult
- using component_le_infnorm[of x]
+ unfolding infnorm_set_image ball_simps inner_scaleR abs_mult
+ using Basis_le_infnorm[of _ x]
apply (auto intro: mult_mono)
done
@@ -2671,9 +2616,13 @@
lemma infnorm_le_norm: "infnorm x \<le> norm x"
unfolding infnorm_def Sup_finite_le_iff[OF infnorm_set_lemma]
unfolding infnorm_set_image ball_simps
- by (metis component_le_norm)
-
-lemma norm_le_infnorm: "norm(x) <= sqrt(real DIM('a)) * infnorm(x::'a::euclidean_space)"
+ by (metis Basis_le_norm)
+
+lemma euclidean_inner: "inner x y = (\<Sum>b\<in>Basis. (x \<bullet> b) * (y \<bullet> b))"
+ by (subst (1 2) euclidean_representation[symmetric, where 'a='a])
+ (simp add: inner_setsum_left inner_setsum_right setsum_cases inner_Basis ac_simps if_distrib)
+
+lemma norm_le_infnorm: "norm(x) <= sqrt DIM('a) * infnorm(x::'a::euclidean_space)"
proof -
let ?d = "DIM('a)"
have "real ?d \<ge> 0" by simp
@@ -2683,13 +2632,14 @@
by (simp add: zero_le_mult_iff infnorm_pos_le)
have th1: "x \<bullet> x \<le> (sqrt (real ?d) * infnorm x)^2"
unfolding power_mult_distrib d2
- unfolding real_of_nat_def apply(subst euclidean_inner)
+ unfolding real_of_nat_def
+ apply(subst euclidean_inner)
apply (subst power2_abs[symmetric])
apply (rule order_trans[OF setsum_bounded[where K="\<bar>infnorm x\<bar>\<twosuperior>"]])
apply (auto simp add: power2_eq_square[symmetric])
apply (subst power2_abs[symmetric])
apply (rule power_mono)
- unfolding infnorm_def Sup_finite_ge_iff[OF infnorm_set_lemma]
+ unfolding infnorm_def Sup_finite_ge_iff[OF infnorm_set_lemma]
unfolding infnorm_set_image bex_simps
apply (rule_tac x=i in bexI)
apply auto
@@ -2872,8 +2822,8 @@
subsection {* An ordering on euclidean spaces that will allow us to talk about intervals *}
class ordered_euclidean_space = ord + euclidean_space +
- assumes eucl_le: "x \<le> y \<longleftrightarrow> (\<forall>i < DIM('a). x $$ i \<le> y $$ i)"
- and eucl_less: "x < y \<longleftrightarrow> (\<forall>i < DIM('a). x $$ i < y $$ i)"
+ assumes eucl_le: "x \<le> y \<longleftrightarrow> (\<forall>i\<in>Basis. x \<bullet> i \<le> y \<bullet> i)"
+ and eucl_less: "x < y \<longleftrightarrow> (\<forall>i\<in>Basis. x \<bullet> i < y \<bullet> i)"
lemma eucl_less_not_refl[simp, intro!]: "\<not> x < (x::'a::ordered_euclidean_space)"
unfolding eucl_less[where 'a='a] by auto
@@ -2889,35 +2839,16 @@
lemma atLeastAtMost_singleton_euclidean[simp]:
fixes a :: "'a::ordered_euclidean_space" shows "{a .. a} = {a}"
- by (force simp: eucl_le[where 'a='a] euclidean_eq[where 'a='a])
-
-lemma basis_real_range: "basis ` {..<1} = {1::real}" by auto
-
-instance real::ordered_euclidean_space
- by default (auto simp add: euclidean_component_def)
-
-lemma Eucl_real_simps[simp]:
- "(x::real) $$ 0 = x"
- "(\<chi>\<chi> i. f i) = ((f 0)::real)"
- "\<And>i. i > 0 \<Longrightarrow> x $$ i = 0"
- defer apply(subst euclidean_eq) apply safe
- unfolding euclidean_lambda_beta'
- unfolding euclidean_component_def apply auto
- done
-
-lemma complex_basis[simp]:
- shows "basis 0 = (1::complex)" and "basis 1 = ii" and "basis (Suc 0) = ii"
- unfolding basis_complex_def by auto
-
-lemma DIM_prod[simp]: "DIM('a \<times> 'b) = DIM('b::euclidean_space) + DIM('a::euclidean_space)"
- (* FIXME: why this orientation? Why not "DIM('a) + DIM('b)" ? *)
- unfolding dimension_prod_def by (rule add_commute)
+ by (force simp: eucl_le[where 'a='a] euclidean_eq_iff[where 'a='a])
+
+instance real :: ordered_euclidean_space
+ by default (auto simp add: Basis_real_def)
instantiation prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space
begin
-definition "x \<le> (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i<DIM('a\<times>'b). x $$ i \<le> y $$ i)"
-definition "x < (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i<DIM('a\<times>'b). x $$ i < y $$ i)"
+definition "x \<le> (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i\<in>Basis. x \<bullet> i \<le> y \<bullet> i)"
+definition "x < (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i\<in>Basis. x \<bullet> i < y \<bullet> i)"
instance
by default (auto simp: less_prod_def less_eq_prod_def)
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Fri Dec 14 15:46:01 2012 +0100
@@ -1,5 +1,5 @@
theory Multivariate_Analysis
-imports Fashoda Extended_Real_Limits
+imports Fashoda Extended_Real_Limits Determinants
begin
end
--- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy Fri Dec 14 15:46:01 2012 +0100
@@ -22,8 +22,11 @@
moreover
{assume H: ?lhs
- have bp: "b \<ge> 0" apply-apply(rule order_trans [OF norm_ge_zero])
- apply(rule H[rule_format, of "basis 0::'a"]) by auto
+ have bp: "b \<ge> 0"
+ apply -
+ apply(rule order_trans [OF norm_ge_zero])
+ apply(rule H[rule_format, of "SOME x::'a. x \<in> Basis"])
+ by (auto intro: SOME_Basis norm_Basis)
{fix x :: "'a"
{assume "x = 0"
then have "norm (f x) \<le> b * norm x" by (simp add: linear_0[OF lf] bp)}
@@ -50,8 +53,8 @@
proof-
{
let ?S = "{norm (f x) |x. norm x = 1}"
- have "norm (f (basis 0)) \<in> ?S" unfolding mem_Collect_eq
- apply(rule_tac x="basis 0" in exI) by auto
+ have "norm (f (SOME i. i \<in> Basis)) \<in> ?S"
+ by (auto intro!: exI[of _ "SOME i. i \<in> Basis"] norm_Basis SOME_Basis)
hence Se: "?S \<noteq> {}" by auto
from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b"
unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def)
@@ -70,8 +73,8 @@
qed
lemma onorm_pos_le: assumes lf: "linear (f::'n::euclidean_space \<Rightarrow> 'm::euclidean_space)" shows "0 <= onorm f"
- using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis 0"]]
- using DIM_positive[where 'a='n] by auto
+ using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "SOME i. i \<in> Basis"]]
+ by (simp add: SOME_Basis)
lemma onorm_eq_0: assumes lf: "linear (f::'a::euclidean_space \<Rightarrow> 'b::euclidean_space)"
shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)"
@@ -87,7 +90,7 @@
proof-
let ?f = "\<lambda>x::'a. (y::'b)"
have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
- apply safe apply(rule_tac x="basis 0" in exI) by auto
+ by (auto simp: SOME_Basis intro!: exI[of _ "SOME i. i \<in> Basis"])
show ?thesis
unfolding onorm_def th
apply (rule Sup_unique) by (simp_all add: setle_def)
--- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Fri Dec 14 15:46:01 2012 +0100
@@ -819,35 +819,40 @@
assumes "2 \<le> DIM('a::euclidean_space)"
shows "path_connected((UNIV::'a::euclidean_space set) - {a})"
proof -
- let ?A = "{x::'a. \<exists>i\<in>{..<DIM('a)}. x $$ i < a $$ i}"
- let ?B = "{x::'a. \<exists>i\<in>{..<DIM('a)}. a $$ i < x $$ i}"
+ let ?A = "{x::'a. \<exists>i\<in>Basis. x \<bullet> i < a \<bullet> i}"
+ let ?B = "{x::'a. \<exists>i\<in>Basis. a \<bullet> i < x \<bullet> i}"
have A: "path_connected ?A"
unfolding Collect_bex_eq
proof (rule path_connected_UNION)
- fix i
- assume "i \<in> {..<DIM('a)}"
- then show "(\<chi>\<chi> i. a $$ i - 1) \<in> {x::'a. x $$ i < a $$ i}" by simp
- show "path_connected {x. x $$ i < a $$ i}" unfolding euclidean_component_def
- by (rule convex_imp_path_connected [OF convex_halfspace_lt])
+ fix i :: 'a
+ assume "i \<in> Basis"
+ then show "(\<Sum>i\<in>Basis. (a \<bullet> i - 1)*\<^sub>R i) \<in> {x::'a. x \<bullet> i < a \<bullet> i}" by simp
+ show "path_connected {x. x \<bullet> i < a \<bullet> i}"
+ using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \<bullet> i"]
+ by (simp add: inner_commute)
qed
have B: "path_connected ?B" unfolding Collect_bex_eq
proof (rule path_connected_UNION)
- fix i
- assume "i \<in> {..<DIM('a)}"
- then show "(\<chi>\<chi> i. a $$ i + 1) \<in> {x::'a. a $$ i < x $$ i}" by simp
- show "path_connected {x. a $$ i < x $$ i}" unfolding euclidean_component_def
- by (rule convex_imp_path_connected [OF convex_halfspace_gt])
+ fix i :: 'a
+ assume "i \<in> Basis"
+ then show "(\<Sum>i\<in>Basis. (a \<bullet> i + 1) *\<^sub>R i) \<in> {x::'a. a \<bullet> i < x \<bullet> i}" by simp
+ show "path_connected {x. a \<bullet> i < x \<bullet> i}"
+ using convex_imp_path_connected [OF convex_halfspace_gt, of "a \<bullet> i" i]
+ by (simp add: inner_commute)
qed
- from assms have "1 < DIM('a)" by auto
- then have "a + basis 0 - basis 1 \<in> ?A \<inter> ?B" by auto
+ obtain S :: "'a set" where "S \<subseteq> Basis" "card S = Suc (Suc 0)"
+ using ex_card[OF assms] by auto
+ then obtain b0 b1 :: 'a where "b0 \<in> Basis" "b1 \<in> Basis" "b0 \<noteq> b1"
+ unfolding card_Suc_eq by auto
+ then have "a + b0 - b1 \<in> ?A \<inter> ?B" by (auto simp: inner_simps inner_Basis)
then have "?A \<inter> ?B \<noteq> {}" by fast
with A B have "path_connected (?A \<union> ?B)"
by (rule path_connected_Un)
- also have "?A \<union> ?B = {x. \<exists>i\<in>{..<DIM('a)}. x $$ i \<noteq> a $$ i}"
+ also have "?A \<union> ?B = {x. \<exists>i\<in>Basis. x \<bullet> i \<noteq> a \<bullet> i}"
unfolding neq_iff bex_disj_distrib Collect_disj_eq ..
also have "\<dots> = {x. x \<noteq> a}"
- unfolding Bex_def euclidean_eq [where 'a='a] by simp
+ unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def)
also have "\<dots> = UNIV - {a}" by auto
finally show ?thesis .
qed
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Dec 14 15:46:01 2012 +0100
@@ -13,9 +13,17 @@
"~~/src/HOL/Library/Countable_Set"
Linear_Algebra
"~~/src/HOL/Library/Glbs"
+ "~~/src/HOL/Library/FuncSet"
Norm_Arith
begin
+lemma countable_PiE:
+ "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (F i)) \<Longrightarrow> countable (PiE I F)"
+ by (induct I arbitrary: F rule: finite_induct) (auto simp: PiE_insert_eq)
+
+lemma countable_rat: "countable \<rat>"
+ unfolding Rats_def by auto
+
subsection {* Topological Basis *}
context topological_space
@@ -593,86 +601,74 @@
lemma ball_empty[intro]: "e \<le> 0 ==> ball x e = {}" by simp
+lemma euclidean_dist_l2:
+ fixes x y :: "'a :: euclidean_space"
+ shows "dist x y = setL2 (\<lambda>i. dist (x \<bullet> i) (y \<bullet> i)) Basis"
+ unfolding dist_norm norm_eq_sqrt_inner setL2_def
+ by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left)
+
+definition "box a b = {x. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}"
+
lemma rational_boxes:
- fixes x :: "'a\<Colon>ordered_euclidean_space"
+ fixes x :: "'a\<Colon>euclidean_space"
assumes "0 < e"
- shows "\<exists>a b. (\<forall>i. a $$ i \<in> \<rat>) \<and> (\<forall>i. b $$ i \<in> \<rat>) \<and> x \<in> {a <..< b} \<and> {a <..< b} \<subseteq> ball x e"
+ shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat> ) \<and> x \<in> box a b \<and> box a b \<subseteq> ball x e"
proof -
def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
- then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos)
- have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x $$ i \<and> x $$ i - y < e'" (is "\<forall>i. ?th i")
+ then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos simp: DIM_positive)
+ have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i")
proof
- fix i from Rats_dense_in_real[of "x $$ i - e'" "x $$ i"] e
- show "?th i" by auto
+ fix i from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e show "?th i" by auto
qed
from choice[OF this] guess a .. note a = this
- have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x $$ i < y \<and> y - x $$ i < e'" (is "\<forall>i. ?th i")
+ have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i")
proof
- fix i from Rats_dense_in_real[of "x $$ i" "x $$ i + e'"] e
- show "?th i" by auto
+ fix i from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e show "?th i" by auto
qed
from choice[OF this] guess b .. note b = this
- { fix y :: 'a assume *: "Chi a < y" "y < Chi b"
- have "dist x y = sqrt (\<Sum>i<DIM('a). (dist (x $$ i) (y $$ i))\<twosuperior>)"
+ let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i"
+ show ?thesis
+ proof (rule exI[of _ ?a], rule exI[of _ ?b], safe)
+ fix y :: 'a assume *: "y \<in> box ?a ?b"
+ have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<twosuperior>)"
unfolding setL2_def[symmetric] by (rule euclidean_dist_l2)
- also have "\<dots> < sqrt (\<Sum>i<DIM('a). e^2 / real (DIM('a)))"
+ also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
proof (rule real_sqrt_less_mono, rule setsum_strict_mono)
- fix i assume i: "i \<in> {..<DIM('a)}"
- have "a i < y$$i \<and> y$$i < b i" using * i eucl_less[where 'a='a] by auto
- moreover have "a i < x$$i" "x$$i - a i < e'" using a by auto
- moreover have "x$$i < b i" "b i - x$$i < e'" using b by auto
- ultimately have "\<bar>x$$i - y$$i\<bar> < 2 * e'" by auto
- then have "dist (x $$ i) (y $$ i) < e/sqrt (real (DIM('a)))"
+ fix i :: "'a" assume i: "i \<in> Basis"
+ have "a i < y\<bullet>i \<and> y\<bullet>i < b i" using * i by (auto simp: box_def)
+ moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" using a by auto
+ moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'" using b by auto
+ ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" by auto
+ then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
unfolding e'_def by (auto simp: dist_real_def)
- then have "(dist (x $$ i) (y $$ i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
+ then have "(dist (x \<bullet> i) (y \<bullet> i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
by (rule power_strict_mono) auto
- then show "(dist (x $$ i) (y $$ i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
+ then show "(dist (x \<bullet> i) (y \<bullet> i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
by (simp add: power_divide)
qed auto
- also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat DIM_positive)
- finally have "dist x y < e" . }
- with a b show ?thesis
- apply (rule_tac exI[of _ "Chi a"])
- apply (rule_tac exI[of _ "Chi b"])
- using eucl_less[where 'a='a] by auto
-qed
-
-lemma ex_rat_list:
- fixes x :: "'a\<Colon>ordered_euclidean_space"
- assumes "\<And> i. x $$ i \<in> \<rat>"
- shows "\<exists> r. length r = DIM('a) \<and> (\<forall> i < DIM('a). of_rat (r ! i) = x $$ i)"
-proof -
- have "\<forall>i. \<exists>r. x $$ i = of_rat r" using assms unfolding Rats_def by blast
- from choice[OF this] guess r ..
- then show ?thesis by (auto intro!: exI[of _ "map r [0 ..< DIM('a)]"])
-qed
-
-lemma open_UNION:
- fixes M :: "'a\<Colon>ordered_euclidean_space set"
- assumes "open M"
- shows "M = UNION {(a, b) | a b. {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)} \<subseteq> M}
- (\<lambda> (a, b). {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)})"
- (is "M = UNION ?idx ?box")
+ also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat)
+ finally show "y \<in> ball x e" by (auto simp: ball_def)
+ qed (insert a b, auto simp: box_def)
+qed
+
+lemma open_UNION_box:
+ fixes M :: "'a\<Colon>euclidean_space set"
+ assumes "open M"
+ defines "a' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)"
+ defines "b' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)"
+ defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^isub>E \<rat> \<times> \<rat>. box (a' f) (b' f) \<subseteq> M}"
+ shows "M = (\<Union>f\<in>I. box (a' f) (b' f))"
proof safe
fix x assume "x \<in> M"
obtain e where e: "e > 0" "ball x e \<subseteq> M"
- using openE[OF assms `x \<in> M`] by auto
- then obtain a b where ab: "x \<in> {a <..< b}" "\<And>i. a $$ i \<in> \<rat>" "\<And>i. b $$ i \<in> \<rat>" "{a <..< b} \<subseteq> ball x e"
- using rational_boxes[OF e(1)] by blast
- then obtain p q where pq: "length p = DIM ('a)"
- "length q = DIM ('a)"
- "\<forall> i < DIM ('a). of_rat (p ! i) = a $$ i \<and> of_rat (q ! i) = b $$ i"
- using ex_rat_list[OF ab(2)] ex_rat_list[OF ab(3)] by blast
- hence p: "Chi (of_rat \<circ> op ! p) = a"
- using euclidean_eq[of "Chi (of_rat \<circ> op ! p)" a]
- unfolding o_def by auto
- from pq have q: "Chi (of_rat \<circ> op ! q) = b"
- using euclidean_eq[of "Chi (of_rat \<circ> op ! q)" b]
- unfolding o_def by auto
- have "x \<in> ?box (p, q)"
- using p q ab by auto
- thus "x \<in> UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto
-qed auto
+ using openE[OF `open M` `x \<in> M`] by auto
+ moreover then obtain a b where ab: "x \<in> box a b"
+ "\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>" "\<forall>i\<in>Basis. b \<bullet> i \<in> \<rat>" "box a b \<subseteq> ball x e"
+ using rational_boxes[OF e(1)] by metis
+ ultimately show "x \<in> (\<Union>f\<in>I. box (a' f) (b' f))"
+ by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"])
+ (auto simp: euclidean_representation I_def a'_def b'_def)
+qed (auto simp: I_def)
subsection{* Connectedness *}
@@ -1156,14 +1152,10 @@
"eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
unfolding eventually_at dist_nz by auto
-lemma eventually_within: "eventually P (at a within S) \<longleftrightarrow>
+lemma eventually_within: (* FIXME: this replaces Limits.eventually_within *)
+ "eventually P (at a within S) \<longleftrightarrow>
(\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
-unfolding eventually_within eventually_at dist_nz by auto
-
-lemma eventually_within_le: "eventually P (at a within S) \<longleftrightarrow>
- (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a <= d \<longrightarrow> P x)" (is "?lhs = ?rhs")
-unfolding eventually_within
-by auto (metis dense order_le_less_trans)
+ by (rule eventually_within_less)
lemma eventually_happens: "eventually P net ==> trivial_limit net \<or> (\<exists>x. P x)"
unfolding trivial_limit_def
@@ -1721,7 +1713,7 @@
assume "\<not> (\<exists>y\<in>A. dist y x < e)"
hence "infdist x A \<ge> e" using `a \<in> A`
unfolding infdist_def
- by (force intro: Inf_greatest simp: dist_commute)
+ by (force simp: dist_commute)
with x `0 < e` show False by auto
qed
qed
@@ -2374,56 +2366,41 @@
by auto
qed
-lemma bounded_component: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $$ i) ` s)"
- apply (erule bounded_linear_image)
- apply (rule bounded_linear_euclidean_component)
- done
-
lemma compact_lemma:
fixes f :: "nat \<Rightarrow> 'a::euclidean_space"
assumes "bounded s" and "\<forall>n. f n \<in> s"
- shows "\<forall>d. \<exists>l::'a. \<exists> r. subseq r \<and>
- (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $$ i) (l $$ i) < e) sequentially)"
-proof
- fix d'::"nat set" def d \<equiv> "d' \<inter> {..<DIM('a)}"
- have "finite d" "d\<subseteq>{..<DIM('a)}" unfolding d_def by auto
- hence "\<exists>l::'a. \<exists>r. subseq r \<and>
- (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $$ i) (l $$ i) < e) sequentially)"
+ shows "\<forall>d\<subseteq>Basis. \<exists>l::'a. \<exists> r. subseq r \<and>
+ (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)"
+proof safe
+ fix d :: "'a set" assume d: "d \<subseteq> Basis"
+ with finite_Basis have "finite d" by (blast intro: finite_subset)
+ from this d show "\<exists>l::'a. \<exists>r. subseq r \<and>
+ (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)"
proof(induct d) case empty thus ?case unfolding subseq_def by auto
- next case (insert k d) have k[intro]:"k<DIM('a)" using insert by auto
- have s': "bounded ((\<lambda>x. x $$ k) ` s)" using `bounded s` by (rule bounded_component)
+ next case (insert k d) have k[intro]:"k\<in>Basis" using insert by auto
+ have s': "bounded ((\<lambda>x. x \<bullet> k) ` s)" using `bounded s`
+ by (auto intro!: bounded_linear_image bounded_linear_inner_left)
obtain l1::"'a" and r1 where r1:"subseq r1" and
- lr1:"\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $$ i) (l1 $$ i) < e) sequentially"
+ lr1:"\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially"
using insert(3) using insert(4) by auto
- have f': "\<forall>n. f (r1 n) $$ k \<in> (\<lambda>x. x $$ k) ` s" using `\<forall>n. f n \<in> s` by simp
- obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) $$ k) ---> l2) sequentially"
+ have f': "\<forall>n. f (r1 n) \<bullet> k \<in> (\<lambda>x. x \<bullet> k) ` s" using `\<forall>n. f n \<in> s` by simp
+ obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) \<bullet> k) ---> l2) sequentially"
using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
def r \<equiv> "r1 \<circ> r2" have r:"subseq r"
using r1 and r2 unfolding r_def o_def subseq_def by auto
moreover
- def l \<equiv> "(\<chi>\<chi> i. if i = k then l2 else l1$$i)::'a"
+ def l \<equiv> "(\<Sum>i\<in>Basis. (if i = k then l2 else l1\<bullet>i) *\<^sub>R i)::'a"
{ fix e::real assume "e>0"
- from lr1 `e>0` have N1:"eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $$ i) (l1 $$ i) < e) sequentially" by blast
- from lr2 `e>0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) $$ k) l2 < e) sequentially" by (rule tendstoD)
- from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) $$ i) (l1 $$ i) < e) sequentially"
+ from lr1 `e>0` have N1:"eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially" by blast
+ from lr2 `e>0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) \<bullet> k) l2 < e) sequentially" by (rule tendstoD)
+ from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) \<bullet> i) (l1 \<bullet> i) < e) sequentially"
by (rule eventually_subseq)
- have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) $$ i) (l $$ i) < e) sequentially"
- using N1' N2 apply(rule eventually_elim2) unfolding l_def r_def o_def
- using insert.prems by auto
+ have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially"
+ using N1' N2
+ by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def)
}
ultimately show ?case by auto
qed
- thus "\<exists>l::'a. \<exists>r. subseq r \<and>
- (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d'. dist (f (r n) $$ i) (l $$ i) < e) sequentially)"
- apply safe apply(rule_tac x=l in exI,rule_tac x=r in exI) apply safe
- apply(erule_tac x=e in allE) unfolding d_def eventually_sequentially apply safe
- apply(rule_tac x=N in exI) apply safe apply(erule_tac x=n in allE,safe)
- apply(erule_tac x=i in ballE)
- proof- fix i and r::"nat=>nat" and n::nat and e::real and l::'a
- assume "i\<in>d'" "i \<notin> d' \<inter> {..<DIM('a)}" and e:"e>0"
- hence *:"i\<ge>DIM('a)" by auto
- thus "dist (f (r n) $$ i) (l $$ i) < e" using e by auto
- qed
qed
instance euclidean_space \<subseteq> heine_borel
@@ -2431,22 +2408,20 @@
fix s :: "'a set" and f :: "nat \<Rightarrow> 'a"
assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
then obtain l::'a and r where r: "subseq r"
- and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>UNIV. dist (f (r n) $$ i) (l $$ i) < e) sequentially"
+ and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially"
using compact_lemma [OF s f] by blast
- let ?d = "{..<DIM('a)}"
{ fix e::real assume "e>0"
- hence "0 < e / (real_of_nat (card ?d))"
- using DIM_positive using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto
- with l have "eventually (\<lambda>n. \<forall>i. dist (f (r n) $$ i) (l $$ i) < e / (real_of_nat (card ?d))) sequentially"
+ hence "0 < e / real_of_nat DIM('a)" by (auto intro!: divide_pos_pos DIM_positive)
+ with l have "eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))) sequentially"
by simp
moreover
- { fix n assume n: "\<forall>i. dist (f (r n) $$ i) (l $$ i) < e / (real_of_nat (card ?d))"
- have "dist (f (r n)) l \<le> (\<Sum>i\<in>?d. dist (f (r n) $$ i) (l $$ i))"
+ { fix n assume n: "\<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))"
+ have "dist (f (r n)) l \<le> (\<Sum>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i))"
apply(subst euclidean_dist_l2) using zero_le_dist by (rule setL2_le_setsum)
- also have "\<dots> < (\<Sum>i\<in>?d. e / (real_of_nat (card ?d)))"
+ also have "\<dots> < (\<Sum>i\<in>(Basis::'a set). e / (real_of_nat DIM('a)))"
apply(rule setsum_strict_mono) using n by auto
- finally have "dist (f (r n)) l < e" unfolding setsum_constant
- using DIM_positive[where 'a='a] by auto
+ finally have "dist (f (r n)) l < e"
+ by auto
}
ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
by (rule eventually_elim1)
@@ -3885,10 +3860,6 @@
shows "continuous F (\<lambda>x. inner (f x) (g x))"
using assms unfolding continuous_def by (rule tendsto_inner)
-lemma continuous_euclidean_component:
- shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x $$ i)"
- unfolding continuous_def by (rule tendsto_euclidean_component)
-
lemma continuous_inverse:
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
assumes "continuous F f" and "f (netlimit F) \<noteq> 0"
@@ -3909,10 +3880,8 @@
lemmas continuous_intros = continuous_at_id continuous_within_id
continuous_const continuous_dist continuous_norm continuous_infnorm
- continuous_add continuous_minus continuous_diff
- continuous_scaleR continuous_mult
- continuous_inner continuous_euclidean_component
- continuous_at_inverse continuous_at_within_inverse
+ continuous_add continuous_minus continuous_diff continuous_scaleR continuous_mult
+ continuous_inner continuous_at_inverse continuous_at_within_inverse
subsubsection {* Structural rules for setwise continuity *}
@@ -3976,11 +3945,6 @@
using bounded_bilinear_inner assms
by (rule bounded_bilinear.continuous_on)
-lemma continuous_on_euclidean_component:
- "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x $$ i)"
- using bounded_linear_euclidean_component
- by (rule bounded_linear.continuous_on)
-
lemma continuous_on_inverse:
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra"
assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0"
@@ -4110,7 +4074,7 @@
continuous_on_compose continuous_on_norm continuous_on_infnorm
continuous_on_add continuous_on_minus continuous_on_diff
continuous_on_scaleR continuous_on_mult continuous_on_inverse
- continuous_on_inner continuous_on_euclidean_component
+ continuous_on_inner
uniformly_continuous_on_id uniformly_continuous_on_const
uniformly_continuous_on_dist uniformly_continuous_on_norm
uniformly_continuous_on_compose uniformly_continuous_on_add
@@ -5070,65 +5034,64 @@
subsection {* Intervals *}
lemma interval: fixes a :: "'a::ordered_euclidean_space" shows
- "{a <..< b} = {x::'a. \<forall>i<DIM('a). a$$i < x$$i \<and> x$$i < b$$i}" and
- "{a .. b} = {x::'a. \<forall>i<DIM('a). a$$i \<le> x$$i \<and> x$$i \<le> b$$i}"
+ "{a <..< b} = {x::'a. \<forall>i\<in>Basis. a\<bullet>i < x\<bullet>i \<and> x\<bullet>i < b\<bullet>i}" and
+ "{a .. b} = {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i}"
by(auto simp add:set_eq_iff eucl_le[where 'a='a] eucl_less[where 'a='a])
lemma mem_interval: fixes a :: "'a::ordered_euclidean_space" shows
- "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i < x$$i \<and> x$$i < b$$i)"
- "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i \<le> x$$i \<and> x$$i \<le> b$$i)"
+ "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < x\<bullet>i \<and> x\<bullet>i < b\<bullet>i)"
+ "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i)"
using interval[of a b] by(auto simp add: set_eq_iff eucl_le[where 'a='a] eucl_less[where 'a='a])
lemma interval_eq_empty: fixes a :: "'a::ordered_euclidean_space" shows
- "({a <..< b} = {} \<longleftrightarrow> (\<exists>i<DIM('a). b$$i \<le> a$$i))" (is ?th1) and
- "({a .. b} = {} \<longleftrightarrow> (\<exists>i<DIM('a). b$$i < a$$i))" (is ?th2)
+ "({a <..< b} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i))" (is ?th1) and
+ "({a .. b} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i < a\<bullet>i))" (is ?th2)
proof-
- { fix i x assume i:"i<DIM('a)" and as:"b$$i \<le> a$$i" and x:"x\<in>{a <..< b}"
- hence "a $$ i < x $$ i \<and> x $$ i < b $$ i" unfolding mem_interval by auto
- hence "a$$i < b$$i" by auto
+ { fix i x assume i:"i\<in>Basis" and as:"b\<bullet>i \<le> a\<bullet>i" and x:"x\<in>{a <..< b}"
+ hence "a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i" unfolding mem_interval by auto
+ hence "a\<bullet>i < b\<bullet>i" by auto
hence False using as by auto }
moreover
- { assume as:"\<forall>i<DIM('a). \<not> (b$$i \<le> a$$i)"
+ { assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i \<le> a\<bullet>i)"
let ?x = "(1/2) *\<^sub>R (a + b)"
- { fix i assume i:"i<DIM('a)"
- have "a$$i < b$$i" using as[THEN spec[where x=i]] using i by auto
- hence "a$$i < ((1/2) *\<^sub>R (a+b)) $$ i" "((1/2) *\<^sub>R (a+b)) $$ i < b$$i"
- unfolding euclidean_simps by auto }
+ { fix i :: 'a assume i:"i\<in>Basis"
+ have "a\<bullet>i < b\<bullet>i" using as[THEN bspec[where x=i]] i by auto
+ hence "a\<bullet>i < ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i < b\<bullet>i"
+ by (auto simp: inner_add_left) }
hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto }
ultimately show ?th1 by blast
- { fix i x assume i:"i<DIM('a)" and as:"b$$i < a$$i" and x:"x\<in>{a .. b}"
- hence "a $$ i \<le> x $$ i \<and> x $$ i \<le> b $$ i" unfolding mem_interval by auto
- hence "a$$i \<le> b$$i" by auto
+ { fix i x assume i:"i\<in>Basis" and as:"b\<bullet>i < a\<bullet>i" and x:"x\<in>{a .. b}"
+ hence "a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" unfolding mem_interval by auto
+ hence "a\<bullet>i \<le> b\<bullet>i" by auto
hence False using as by auto }
moreover
- { assume as:"\<forall>i<DIM('a). \<not> (b$$i < a$$i)"
+ { assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)"
let ?x = "(1/2) *\<^sub>R (a + b)"
- { fix i assume i:"i<DIM('a)"
- have "a$$i \<le> b$$i" using as[THEN spec[where x=i]] by auto
- hence "a$$i \<le> ((1/2) *\<^sub>R (a+b)) $$ i" "((1/2) *\<^sub>R (a+b)) $$ i \<le> b$$i"
- unfolding euclidean_simps by auto }
+ { fix i :: 'a assume i:"i\<in>Basis"
+ have "a\<bullet>i \<le> b\<bullet>i" using as[THEN bspec[where x=i]] i by auto
+ hence "a\<bullet>i \<le> ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i \<le> b\<bullet>i"
+ by (auto simp: inner_add_left) }
hence "{a .. b} \<noteq> {}" using mem_interval(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} \<noteq> {} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i \<le> b$$i)" and
- "{a <..< b} \<noteq> {} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i < b$$i)"
+ "{a .. b} \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i)" and
+ "{a <..< b} \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)"
unfolding interval_eq_empty[of a b] by fastforce+
lemma interval_sing:
fixes a :: "'a::ordered_euclidean_space"
shows "{a .. a} = {a}" and "{a<..<a} = {}"
unfolding set_eq_iff mem_interval eq_iff [symmetric]
- by (auto simp add: euclidean_eq[where 'a='a] eq_commute
- eucl_less[where 'a='a] eucl_le[where 'a='a])
+ by (auto intro: euclidean_eqI simp: ex_in_conv)
lemma subset_interval_imp: fixes a :: "'a::ordered_euclidean_space" shows
- "(\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
- "(\<forall>i<DIM('a). a$$i < c$$i \<and> d$$i < b$$i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
- "(\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
- "(\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i) \<Longrightarrow> {c<..<d} \<subseteq> {a<..<b}"
+ "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
+ "(\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
+ "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
+ "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> {c<..<d} \<subseteq> {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)+
@@ -5139,94 +5102,96 @@
by (fast intro: less_imp_le)
lemma subset_interval: fixes a :: "'a::ordered_euclidean_space" shows
- "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i<DIM('a). c$$i \<le> d$$i) --> (\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i)" (is ?th1) and
- "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i<DIM('a). c$$i \<le> d$$i) --> (\<forall>i<DIM('a). a$$i < c$$i \<and> d$$i < b$$i)" (is ?th2) and
- "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i<DIM('a). c$$i < d$$i) --> (\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i)" (is ?th3) and
- "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i<DIM('a). c$$i < d$$i) --> (\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i)" (is ?th4)
+ "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th1) and
+ "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i)" (is ?th2) and
+ "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th3) and
+ "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th4)
proof-
show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans)
show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le)
- { assume as: "{c<..<d} \<subseteq> {a .. b}" "\<forall>i<DIM('a). c$$i < d$$i"
+ { assume as: "{c<..<d} \<subseteq> {a .. b}" "\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i"
hence "{c<..<d} \<noteq> {}" unfolding interval_eq_empty by auto
- fix i assume i:"i<DIM('a)"
+ fix i :: 'a assume i:"i\<in>Basis"
(** TODO combine the following two parts as done in the HOL_light version. **)
- { let ?x = "(\<chi>\<chi> j. (if j=i then ((min (a$$j) (d$$j))+c$$j)/2 else (c$$j+d$$j)/2))::'a"
- assume as2: "a$$i > c$$i"
- { fix j assume j:"j<DIM('a)"
- hence "c $$ j < ?x $$ j \<and> ?x $$ j < d $$ j"
- apply(cases "j=i") using as(2)[THEN spec[where x=j]] i
+ { let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((min (a\<bullet>j) (d\<bullet>j))+c\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a"
+ assume as2: "a\<bullet>i > c\<bullet>i"
+ { fix j :: 'a assume j:"j\<in>Basis"
+ hence "c \<bullet> j < ?x \<bullet> j \<and> ?x \<bullet> j < d \<bullet> j"
+ apply(cases "j=i") using as(2)[THEN bspec[where x=j]] i
by (auto simp add: as2) }
hence "?x\<in>{c<..<d}" using i unfolding mem_interval by auto
moreover
have "?x\<notin>{a .. b}"
- unfolding mem_interval apply auto apply(rule_tac x=i in exI)
- using as(2)[THEN spec[where x=i]] and as2 i
+ unfolding mem_interval apply auto apply(rule_tac x=i in bexI)
+ using as(2)[THEN bspec[where x=i]] and as2 i
by auto
ultimately have False using as by auto }
- hence "a$$i \<le> c$$i" by(rule ccontr)auto
+ hence "a\<bullet>i \<le> c\<bullet>i" by(rule ccontr)auto
moreover
- { let ?x = "(\<chi>\<chi> j. (if j=i then ((max (b$$j) (c$$j))+d$$j)/2 else (c$$j+d$$j)/2))::'a"
- assume as2: "b$$i < d$$i"
- { fix j assume "j<DIM('a)"
- hence "d $$ j > ?x $$ j \<and> ?x $$ j > c $$ j"
- apply(cases "j=i") using as(2)[THEN spec[where x=j]]
- by (auto simp add: as2) }
+ { let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((max (b\<bullet>j) (c\<bullet>j))+d\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a"
+ assume as2: "b\<bullet>i < d\<bullet>i"
+ { fix j :: 'a assume "j\<in>Basis"
+ hence "d \<bullet> j > ?x \<bullet> j \<and> ?x \<bullet> j > c \<bullet> j"
+ apply(cases "j=i") using as(2)[THEN bspec[where x=j]]
+ by (auto simp add: as2) }
hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
moreover
have "?x\<notin>{a .. b}"
- unfolding mem_interval apply auto apply(rule_tac x=i in exI)
- using as(2)[THEN spec[where x=i]] and as2 using i
+ unfolding mem_interval apply auto apply(rule_tac x=i in bexI)
+ using as(2)[THEN bspec[where x=i]] and as2 using i
by auto
ultimately have False using as by auto }
- hence "b$$i \<ge> d$$i" by(rule ccontr)auto
+ hence "b\<bullet>i \<ge> d\<bullet>i" by(rule ccontr)auto
ultimately
- have "a$$i \<le> c$$i \<and> d$$i \<le> b$$i" by auto
+ have "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" by auto
} note part1 = this
- show ?th3 unfolding subset_eq and Ball_def and mem_interval
- apply(rule,rule,rule,rule) apply(rule part1) unfolding subset_eq and Ball_def and mem_interval
- prefer 4 apply auto by(erule_tac x=i in allE,erule_tac x=i in allE,fastforce)+
- { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i<DIM('a). c$$i < d$$i"
- fix i assume i:"i<DIM('a)"
+ show ?th3
+ unfolding subset_eq and Ball_def and mem_interval
+ apply(rule,rule,rule,rule)
+ apply(rule part1)
+ unfolding subset_eq and Ball_def and mem_interval
+ prefer 4
+ apply auto
+ by(erule_tac x=xa in allE,erule_tac x=xa in allE,fastforce)+
+ { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i"
+ fix i :: 'a assume i:"i\<in>Basis"
from as(1) have "{c<..<d} \<subseteq> {a..b}" using interval_open_subset_closed[of a b] by auto
- hence "a$$i \<le> c$$i \<and> d$$i \<le> b$$i" using part1 and as(2) using i by auto } note * = this
+ hence "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" using part1 and as(2) using i by auto } note * = this
show ?th4 unfolding subset_eq and Ball_def and mem_interval
apply(rule,rule,rule,rule) apply(rule *) unfolding subset_eq and Ball_def and mem_interval prefer 4
- apply auto by(erule_tac x=i in allE, simp)+
-qed
-
-lemma disjoint_interval: fixes a::"'a::ordered_euclidean_space" shows
- "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i<DIM('a). (b$$i < a$$i \<or> d$$i < c$$i \<or> b$$i < c$$i \<or> d$$i < a$$i))" (is ?th1) and
- "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i<DIM('a). (b$$i < a$$i \<or> d$$i \<le> c$$i \<or> b$$i \<le> c$$i \<or> d$$i \<le> a$$i))" (is ?th2) and
- "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i<DIM('a). (b$$i \<le> a$$i \<or> d$$i < c$$i \<or> b$$i \<le> c$$i \<or> d$$i \<le> a$$i))" (is ?th3) and
- "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i<DIM('a). (b$$i \<le> a$$i \<or> d$$i \<le> c$$i \<or> b$$i \<le> c$$i \<or> d$$i \<le> a$$i))" (is ?th4)
-proof-
- let ?z = "(\<chi>\<chi> i. ((max (a$$i) (c$$i)) + (min (b$$i) (d$$i))) / 2)::'a"
- note * = set_eq_iff Int_iff empty_iff mem_interval all_conj_distrib[THEN sym] eq_False
- show ?th1 unfolding * apply safe apply(erule_tac x="?z" in allE)
- unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto
- show ?th2 unfolding * apply safe apply(erule_tac x="?z" in allE)
- unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto
- show ?th3 unfolding * apply safe apply(erule_tac x="?z" in allE)
- unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto
- show ?th4 unfolding * apply safe apply(erule_tac x="?z" in allE)
- unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto
+ apply auto by(erule_tac x=xa in allE, simp)+
qed
lemma inter_interval: fixes a :: "'a::ordered_euclidean_space" shows
- "{a .. b} \<inter> {c .. d} = {(\<chi>\<chi> i. max (a$$i) (c$$i)) .. (\<chi>\<chi> i. min (b$$i) (d$$i))}"
- unfolding set_eq_iff and Int_iff and mem_interval
- by auto
+ "{a .. b} \<inter> {c .. d} = {(\<Sum>i\<in>Basis. max (a\<bullet>i) (c\<bullet>i) *\<^sub>R i) .. (\<Sum>i\<in>Basis. min (b\<bullet>i) (d\<bullet>i) *\<^sub>R i)}"
+ unfolding set_eq_iff and Int_iff and mem_interval by auto
+
+lemma disjoint_interval: fixes a::"'a::ordered_euclidean_space" shows
+ "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i < c\<bullet>i \<or> d\<bullet>i < a\<bullet>i))" (is ?th1) and
+ "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th2) and
+ "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th3) and
+ "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th4)
+proof-
+ let ?z = "(\<Sum>i\<in>Basis. (((max (a\<bullet>i) (c\<bullet>i)) + (min (b\<bullet>i) (d\<bullet>i))) / 2) *\<^sub>R i)::'a"
+ have **: "\<And>P Q. (\<And>i :: 'a. i \<in> Basis \<Longrightarrow> Q ?z i \<Longrightarrow> P i) \<Longrightarrow>
+ (\<And>i x :: 'a. i \<in> Basis \<Longrightarrow> P i \<Longrightarrow> Q x i) \<Longrightarrow> (\<forall>x. \<exists>i\<in>Basis. Q x i) \<longleftrightarrow> (\<exists>i\<in>Basis. P i)"
+ by blast
+ note * = set_eq_iff Int_iff empty_iff mem_interval 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 *)
lemma open_interval[intro]:
fixes a b :: "'a::ordered_euclidean_space" shows "open {a<..<b}"
proof-
- have "open (\<Inter>i<DIM('a). (\<lambda>x. x$$i) -` {a$$i<..<b$$i})"
+ have "open (\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i<..<b\<bullet>i})"
by (intro open_INT finite_lessThan ballI continuous_open_vimage allI
- linear_continuous_at bounded_linear_euclidean_component
- open_real_greaterThanLessThan)
- also have "(\<Inter>i<DIM('a). (\<lambda>x. x$$i) -` {a$$i<..<b$$i}) = {a<..<b}"
+ linear_continuous_at open_real_greaterThanLessThan finite_Basis bounded_linear_inner_left)
+ also have "(\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i<..<b\<bullet>i}) = {a<..<b}"
by (auto simp add: eucl_less [where 'a='a])
finally show "open {a<..<b}" .
qed
@@ -5234,11 +5199,10 @@
lemma closed_interval[intro]:
fixes a b :: "'a::ordered_euclidean_space" shows "closed {a .. b}"
proof-
- have "closed (\<Inter>i<DIM('a). (\<lambda>x. x$$i) -` {a$$i .. b$$i})"
+ have "closed (\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i})"
by (intro closed_INT ballI continuous_closed_vimage allI
- linear_continuous_at bounded_linear_euclidean_component
- closed_real_atLeastAtMost)
- also have "(\<Inter>i<DIM('a). (\<lambda>x. x$$i) -` {a$$i .. b$$i}) = {a .. b}"
+ linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left)
+ also have "(\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i}) = {a .. b}"
by (auto simp add: eucl_le [where 'a='a])
finally show "closed {a .. b}" .
qed
@@ -5253,29 +5217,29 @@
{ fix x assume "x \<in> interior {a..b}"
then obtain s where s:"open s" "x \<in> s" "s \<subseteq> {a..b}" ..
then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a..b}" unfolding open_dist and subset_eq by auto
- { fix i assume i:"i<DIM('a)"
- have "dist (x - (e / 2) *\<^sub>R basis i) x < e"
- "dist (x + (e / 2) *\<^sub>R basis i) x < e"
+ { fix i :: 'a assume i:"i\<in>Basis"
+ have "dist (x - (e / 2) *\<^sub>R i) x < e"
+ "dist (x + (e / 2) *\<^sub>R i) x < e"
unfolding dist_norm apply auto
- unfolding norm_minus_cancel using norm_basis and `e>0` by auto
- hence "a $$ i \<le> (x - (e / 2) *\<^sub>R basis i) $$ i"
- "(x + (e / 2) *\<^sub>R basis i) $$ i \<le> b $$ i"
- using e[THEN spec[where x="x - (e/2) *\<^sub>R basis i"]]
- and e[THEN spec[where x="x + (e/2) *\<^sub>R basis i"]]
+ unfolding norm_minus_cancel using norm_Basis[OF i] `e>0` by auto
+ hence "a \<bullet> i \<le> (x - (e / 2) *\<^sub>R i) \<bullet> i"
+ "(x + (e / 2) *\<^sub>R i) \<bullet> i \<le> b \<bullet> 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 using i by blast+
- hence "a $$ i < x $$ i" and "x $$ i < b $$ i" unfolding euclidean_simps
- unfolding basis_component using `e>0` i by auto }
+ hence "a \<bullet> i < x \<bullet> i" and "x \<bullet> i < b \<bullet> i"
+ using `e>0` i by (auto simp: inner_diff_left inner_Basis inner_add_left) }
hence "x \<in> {a<..<b}" unfolding mem_interval by auto }
thus "?L \<subseteq> ?R" ..
qed
lemma bounded_closed_interval: fixes a :: "'a::ordered_euclidean_space" shows "bounded {a .. b}"
proof-
- let ?b = "\<Sum>i<DIM('a). \<bar>a$$i\<bar> + \<bar>b$$i\<bar>"
- { fix x::"'a" assume x:"\<forall>i<DIM('a). a $$ i \<le> x $$ i \<and> x $$ i \<le> b $$ i"
- { fix i assume "i<DIM('a)"
- hence "\<bar>x$$i\<bar> \<le> \<bar>a$$i\<bar> + \<bar>b$$i\<bar>" using x[THEN spec[where x=i]] by auto }
- hence "(\<Sum>i<DIM('a). \<bar>x $$ i\<bar>) \<le> ?b" apply-apply(rule setsum_mono) by auto
+ let ?b = "\<Sum>i\<in>Basis. \<bar>a\<bullet>i\<bar> + \<bar>b\<bullet>i\<bar>"
+ { fix x::"'a" assume x:"\<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i"
+ { fix i :: 'a assume "i\<in>Basis"
+ hence "\<bar>x\<bullet>i\<bar> \<le> \<bar>a\<bullet>i\<bar> + \<bar>b\<bullet>i\<bar>" using x[THEN bspec[where x=i]] by auto }
+ hence "(\<Sum>i\<in>Basis. \<bar>x \<bullet> i\<bar>) \<le> ?b" apply-apply(rule setsum_mono) by auto
hence "norm x \<le> ?b" using norm_le_l1[of x] by auto }
thus ?thesis unfolding interval and bounded_iff by auto
qed
@@ -5298,10 +5262,9 @@
lemma open_interval_midpoint: fixes a :: "'a::ordered_euclidean_space"
assumes "{a<..<b} \<noteq> {}" shows "((1/2) *\<^sub>R (a + b)) \<in> {a<..<b}"
proof-
- { fix i assume "i<DIM('a)"
- hence "a $$ i < ((1 / 2) *\<^sub>R (a + b)) $$ i \<and> ((1 / 2) *\<^sub>R (a + b)) $$ i < b $$ i"
- using assms[unfolded interval_ne_empty, THEN spec[where x=i]]
- unfolding euclidean_simps by auto }
+ { fix i :: 'a assume "i\<in>Basis"
+ hence "a \<bullet> i < ((1 / 2) *\<^sub>R (a + b)) \<bullet> i \<and> ((1 / 2) *\<^sub>R (a + b)) \<bullet> i < b \<bullet> i"
+ using assms[unfolded interval_ne_empty, THEN bspec[where x=i]] by (auto simp: inner_add_left) }
thus ?thesis unfolding mem_interval by auto
qed
@@ -5309,23 +5272,23 @@
assumes x:"x \<in> {a<..<b}" and y:"y \<in> {a .. b}" and e:"0 < e" "e \<le> 1"
shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> {a<..<b}"
proof-
- { fix i assume i:"i<DIM('a)"
- have "a $$ i = e * a$$i + (1 - e) * a$$i" unfolding left_diff_distrib by simp
- also have "\<dots> < e * x $$ i + (1 - e) * y $$ i" apply(rule add_less_le_mono)
+ { fix i :: 'a assume i:"i\<in>Basis"
+ have "a \<bullet> i = e * (a \<bullet> i) + (1 - e) * (a \<bullet> i)" unfolding left_diff_distrib by simp
+ also have "\<dots> < e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" 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 apply simp
using y unfolding mem_interval using i apply simp
done
- finally have "a $$ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i" unfolding euclidean_simps by auto
+ finally have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i" unfolding inner_simps by auto
moreover {
- have "b $$ i = e * b$$i + (1 - e) * b$$i" unfolding left_diff_distrib by simp
- also have "\<dots> > e * x $$ i + (1 - e) * y $$ i" apply(rule add_less_le_mono)
+ have "b \<bullet> i = e * (b\<bullet>i) + (1 - e) * (b\<bullet>i)" unfolding left_diff_distrib by simp
+ also have "\<dots> > e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" 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 apply simp
using y unfolding mem_interval using i apply simp
done
- finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i < b $$ i" unfolding euclidean_simps by auto
- } ultimately have "a $$ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i < b $$ i" by auto }
+ finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" unfolding inner_simps by auto
+ } ultimately have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" by auto }
thus ?thesis unfolding mem_interval by auto
qed
@@ -5365,11 +5328,11 @@
assumes "bounded s" shows "\<exists>a. s \<subseteq> {-a<..<a}"
proof-
obtain b where "b>0" and b:"\<forall>x\<in>s. norm x \<le> b" using assms[unfolded bounded_pos] by auto
- def a \<equiv> "(\<chi>\<chi> i. b+1)::'a"
+ def a \<equiv> "(\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)::'a"
{ fix x assume "x\<in>s"
- fix i assume i:"i<DIM('a)"
- hence "(-a)$$i < x$$i" and "x$$i < a$$i" using b[THEN bspec[where x=x], OF `x\<in>s`]
- and component_le_norm[of x i] unfolding euclidean_simps and a_def by auto }
+ fix i :: 'a assume i:"i\<in>Basis"
+ hence "(-a)\<bullet>i < x\<bullet>i" and "x\<bullet>i < a\<bullet>i" using b[THEN bspec[where x=x], OF `x\<in>s`]
+ and Basis_le_norm[OF i, of x] unfolding inner_simps and a_def by auto }
thus ?thesis by(auto intro: exI[where x=a] simp add: eucl_less[where 'a='a])
qed
@@ -5413,59 +5376,61 @@
(* Some stuff for half-infinite intervals too; FIXME: notation? *)
lemma closed_interval_left: fixes b::"'a::euclidean_space"
- shows "closed {x::'a. \<forall>i<DIM('a). x$$i \<le> b$$i}"
+ shows "closed {x::'a. \<forall>i\<in>Basis. x\<bullet>i \<le> b\<bullet>i}"
proof-
- { fix i assume i:"i<DIM('a)"
- fix x::"'a" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i<DIM('a). x $$ i \<le> b $$ i}. x' \<noteq> x \<and> dist x' x < e"
- { assume "x$$i > b$$i"
- then obtain y where "y $$ i \<le> b $$ i" "y \<noteq> x" "dist y x < x$$i - b$$i"
- using x[THEN spec[where x="x$$i - b$$i"]] using i by auto
- hence False using component_le_norm[of "y - x" i] unfolding dist_norm euclidean_simps using i
- by auto }
- hence "x$$i \<le> b$$i" by(rule ccontr)auto }
+ { fix i :: 'a assume i:"i\<in>Basis"
+ fix x::"'a" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i\<in>Basis. x \<bullet> i \<le> b \<bullet> i}. x' \<noteq> x \<and> dist x' x < e"
+ { assume "x\<bullet>i > b\<bullet>i"
+ then obtain y where "y \<bullet> i \<le> b \<bullet> i" "y \<noteq> x" "dist y x < x\<bullet>i - b\<bullet>i"
+ using x[THEN spec[where x="x\<bullet>i - b\<bullet>i"]] using i by auto
+ hence False using Basis_le_norm[OF i, of "y - x"] unfolding dist_norm inner_simps using i
+ by auto }
+ hence "x\<bullet>i \<le> b\<bullet>i" by(rule ccontr)auto }
thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
qed
lemma closed_interval_right: fixes a::"'a::euclidean_space"
- shows "closed {x::'a. \<forall>i<DIM('a). a$$i \<le> x$$i}"
+ shows "closed {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i}"
proof-
- { fix i assume i:"i<DIM('a)"
- fix x::"'a" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i<DIM('a). a $$ i \<le> x $$ i}. x' \<noteq> x \<and> dist x' x < e"
- { assume "a$$i > x$$i"
- then obtain y where "a $$ i \<le> y $$ i" "y \<noteq> x" "dist y x < a$$i - x$$i"
- using x[THEN spec[where x="a$$i - x$$i"]] i by auto
- hence False using component_le_norm[of "y - x" i] unfolding dist_norm and euclidean_simps by auto }
- hence "a$$i \<le> x$$i" by(rule ccontr)auto }
+ { fix i :: 'a assume i:"i\<in>Basis"
+ fix x::"'a" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i}. x' \<noteq> x \<and> dist x' x < e"
+ { assume "a\<bullet>i > x\<bullet>i"
+ then obtain y where "a \<bullet> i \<le> y \<bullet> i" "y \<noteq> x" "dist y x < a\<bullet>i - x\<bullet>i"
+ using x[THEN spec[where x="a\<bullet>i - x\<bullet>i"]] i by auto
+ hence False using Basis_le_norm[OF i, of "y - x"] unfolding dist_norm inner_simps by auto }
+ hence "a\<bullet>i \<le> x\<bullet>i" by(rule ccontr)auto }
thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
qed
-instance ordered_euclidean_space \<subseteq> countable_basis_space
+lemma open_box: "open (box a b)"
+proof -
+ have "open (\<Inter>i\<in>Basis. (op \<bullet> i) -` {a \<bullet> i <..< b \<bullet> i})"
+ by (auto intro!: continuous_open_vimage continuous_inner continuous_at_id continuous_const)
+ also have "(\<Inter>i\<in>Basis. (op \<bullet> i) -` {a \<bullet> i <..< b \<bullet> i}) = box a b"
+ by (auto simp add: box_def inner_commute)
+ finally show ?thesis .
+qed
+
+instance euclidean_space \<subseteq> countable_basis_space
proof
- def to_cube \<equiv> "\<lambda>(a, b). {Chi (real_of_rat \<circ> op ! a)<..<Chi (real_of_rat \<circ> op ! b)}::'a set"
- def B \<equiv> "(\<lambda>n. (to_cube (from_nat n)::'a set)) ` UNIV"
- have "countable B" unfolding B_def by simp
+ def a \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. fst (f i) *\<^sub>R i"
+ then have a: "\<And>f. (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i) = a f" by simp
+ def b \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. snd (f i) *\<^sub>R i"
+ then have b: "\<And>f. (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i) = b f" by simp
+ def B \<equiv> "(\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^isub>E (\<rat> \<times> \<rat>))"
+
+ have "countable B" unfolding B_def
+ by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat)
moreover
- have "Ball B open" unfolding B_def
+ have "Ball B open" by (simp add: B_def open_box)
+ moreover have "(\<forall>A. open A \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = A))"
proof safe
- fix n show "open (to_cube (from_nat n))"
- by (cases "from_nat n::rat list \<times> rat list")
- (simp add: open_interval to_cube_def)
- qed
- moreover have "(\<forall>x. open x \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = x))"
- proof safe
- fix x::"'a set" assume "open x"
- def lists \<equiv> "{(a, b) |a b. to_cube (a, b) \<subseteq> x}"
- from open_UNION[OF `open x`]
- have "\<Union>(to_cube ` lists) = x" unfolding lists_def to_cube_def
- by simp
- moreover have "to_cube ` lists \<subseteq> B"
- proof
- fix x assume "x \<in> to_cube ` lists"
- then obtain l where "l \<in> lists" "x = to_cube l" by auto
- thus "x \<in> B" by (auto simp add: B_def intro!: image_eqI[where x="to_nat l"])
- qed
- ultimately
- show "\<exists>B'\<subseteq>B. \<Union>B' = x" by blast
+ fix A::"'a set" assume "open A"
+ show "\<exists>B'\<subseteq>B. \<Union>B' = A"
+ apply (rule exI[of _ "{b\<in>B. b \<subseteq> A}"])
+ apply (subst (3) open_UNION_box[OF `open A`])
+ apply (auto simp add: a b B_def)
+ done
qed
ultimately
show "\<exists>B::'a set set. countable B \<and> topological_basis B" unfolding topological_basis_def by blast
@@ -5476,7 +5441,7 @@
text {* Intervals in general, including infinite and mixtures of open and closed. *}
definition "is_interval (s::('a::euclidean_space) set) \<longleftrightarrow>
- (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i<DIM('a). ((a$$i \<le> x$$i \<and> x$$i \<le> b$$i) \<or> (b$$i \<le> x$$i \<and> x$$i \<le> a$$i))) \<longrightarrow> x \<in> s)"
+ (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i\<in>Basis. ((a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i) \<or> (b\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> a\<bullet>i))) \<longrightarrow> x \<in> s)"
lemma is_interval_interval: "is_interval {a .. b::'a::ordered_euclidean_space}" (is ?th1)
"is_interval {a<..<b}" (is ?th2) proof -
@@ -5560,9 +5525,6 @@
lemma continuous_at_inner: "continuous (at x) (inner a)"
unfolding continuous_at by (intro tendsto_intros)
-lemma continuous_at_euclidean_component[intro!, simp]: "continuous (at x) (\<lambda>x. x $$ i)"
- unfolding euclidean_component_def by (rule continuous_at_inner)
-
lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
by (simp add: closed_Collect_le)
@@ -5573,11 +5535,11 @@
by (simp add: closed_Collect_eq)
lemma closed_halfspace_component_le:
- shows "closed {x::'a::euclidean_space. x$$i \<le> a}"
+ shows "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}"
by (simp add: closed_Collect_le)
lemma closed_halfspace_component_ge:
- shows "closed {x::'a::euclidean_space. x$$i \<ge> a}"
+ shows "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}"
by (simp add: closed_Collect_le)
text {* Openness of halfspaces. *}
@@ -5589,33 +5551,33 @@
by (simp add: open_Collect_less)
lemma open_halfspace_component_lt:
- shows "open {x::'a::euclidean_space. x$$i < a}"
+ shows "open {x::'a::euclidean_space. x\<bullet>i < a}"
by (simp add: open_Collect_less)
lemma open_halfspace_component_gt:
- shows "open {x::'a::euclidean_space. x$$i > a}"
+ shows "open {x::'a::euclidean_space. x\<bullet>i > a}"
by (simp add: open_Collect_less)
text{* Instantiation for intervals on @{text ordered_euclidean_space} *}
lemma eucl_lessThan_eq_halfspaces:
fixes a :: "'a\<Colon>ordered_euclidean_space"
- shows "{..<a} = (\<Inter>i<DIM('a). {x. x $$ i < a $$ i})"
+ shows "{..<a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})"
by (auto simp: eucl_less[where 'a='a])
lemma eucl_greaterThan_eq_halfspaces:
fixes a :: "'a\<Colon>ordered_euclidean_space"
- shows "{a<..} = (\<Inter>i<DIM('a). {x. a $$ i < x $$ i})"
+ shows "{a<..} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})"
by (auto simp: eucl_less[where 'a='a])
lemma eucl_atMost_eq_halfspaces:
fixes a :: "'a\<Colon>ordered_euclidean_space"
- shows "{.. a} = (\<Inter>i<DIM('a). {x. x $$ i \<le> a $$ i})"
+ shows "{.. a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i \<le> a \<bullet> i})"
by (auto simp: eucl_le[where 'a='a])
lemma eucl_atLeast_eq_halfspaces:
fixes a :: "'a\<Colon>ordered_euclidean_space"
- shows "{a ..} = (\<Inter>i<DIM('a). {x. a $$ i \<le> x $$ i})"
+ shows "{a ..} = (\<Inter>i\<in>Basis. {x. a \<bullet> i \<le> x \<bullet> i})"
by (auto simp: eucl_le[where 'a='a])
lemma open_eucl_lessThan[simp, intro]:
@@ -5640,35 +5602,24 @@
unfolding eucl_atLeast_eq_halfspaces
by (simp add: closed_INT closed_Collect_le)
-lemma open_vimage_euclidean_component: "open S \<Longrightarrow> open ((\<lambda>x. x $$ i) -` S)"
- by (auto intro!: continuous_open_vimage)
-
text {* This gives a simple derivation of limit component bounds. *}
lemma Lim_component_le: fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
- assumes "(f ---> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. f(x)$$i \<le> b) net"
- shows "l$$i \<le> b"
-proof-
- { fix x have "x \<in> {x::'b. inner (basis i) x \<le> b} \<longleftrightarrow> x$$i \<le> b"
- unfolding euclidean_component_def by auto } note * = this
- show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \<le> b}" f net l] unfolding *
- using closed_halfspace_le[of "(basis i)::'b" b] and assms(1,2,3) by auto
-qed
+ assumes "(f ---> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. f(x)\<bullet>i \<le> b) net"
+ shows "l\<bullet>i \<le> b"
+ by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)])
lemma Lim_component_ge: fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
- assumes "(f ---> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. b \<le> (f x)$$i) net"
- shows "b \<le> l$$i"
-proof-
- { fix x have "x \<in> {x::'b. inner (basis i) x \<ge> b} \<longleftrightarrow> x$$i \<ge> b"
- unfolding euclidean_component_def by auto } note * = this
- show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \<ge> b}" f net l] unfolding *
- using closed_halfspace_ge[of b "(basis i)"] and assms(1,2,3) by auto
-qed
+ assumes "(f ---> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. b \<le> (f x)\<bullet>i) net"
+ shows "b \<le> l\<bullet>i"
+ by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)])
lemma Lim_component_eq: fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
- assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>x. f(x)$$i = b) net"
- shows "l$$i = b"
- using ev[unfolded order_eq_iff eventually_conj_iff] using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto
+ assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>x. f(x)\<bullet>i = b) net"
+ shows "l\<bullet>i = b"
+ using ev[unfolded order_eq_iff eventually_conj_iff]
+ using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto
+
text{* Limits relative to a union. *}
lemma eventually_within_Un:
@@ -5726,9 +5677,8 @@
qed
lemma connected_ivt_component: fixes x::"'a::euclidean_space" shows
- "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$$k \<le> a \<Longrightarrow> a \<le> y$$k \<Longrightarrow> (\<exists>z\<in>s. z$$k = a)"
- using connected_ivt_hyperplane[of s x y "(basis k)::'a" a]
- unfolding euclidean_component_def by auto
+ "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x\<bullet>k \<le> a \<Longrightarrow> a \<le> y\<bullet>k \<Longrightarrow> (\<exists>z\<in>s. z\<bullet>k = a)"
+ using connected_ivt_hyperplane[of s x y "k::'a" a] by (auto simp: inner_commute)
subsection {* Homeomorphisms *}
@@ -5998,95 +5948,91 @@
subsection {* Some properties of a canonical subspace *}
lemma subspace_substandard:
- "subspace {x::'a::euclidean_space. (\<forall>i<DIM('a). P i \<longrightarrow> x$$i = 0)}"
- unfolding subspace_def by auto
+ "subspace {x::'a::euclidean_space. (\<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0)}"
+ unfolding subspace_def by (auto simp: inner_add_left)
lemma closed_substandard:
- "closed {x::'a::euclidean_space. \<forall>i<DIM('a). P i --> x$$i = 0}" (is "closed ?A")
+ "closed {x::'a::euclidean_space. \<forall>i\<in>Basis. P i --> x\<bullet>i = 0}" (is "closed ?A")
proof-
- let ?D = "{i. P i} \<inter> {..<DIM('a)}"
- have "closed (\<Inter>i\<in>?D. {x::'a. x$$i = 0})"
+ let ?D = "{i\<in>Basis. P i}"
+ have "closed (\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0})"
by (simp add: closed_INT closed_Collect_eq)
- also have "(\<Inter>i\<in>?D. {x::'a. x$$i = 0}) = ?A"
+ also have "(\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0}) = ?A"
by auto
finally show "closed ?A" .
qed
-lemma dim_substandard: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
- shows "dim {x::'a::euclidean_space. \<forall>i<DIM('a). i \<notin> d \<longrightarrow> x$$i = 0} = card d" (is "dim ?A = _")
+lemma dim_substandard: assumes d: "d \<subseteq> Basis"
+ shows "dim {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0} = card d" (is "dim ?A = _")
proof-
- let ?D = "{..<DIM('a)}"
- let ?B = "(basis::nat => 'a) ` d"
- let ?bas = "basis::nat \<Rightarrow> 'a"
- have "?B \<subseteq> ?A" by auto
+ let ?D = "Basis :: 'a set"
+ have "d \<subseteq> ?A" using d by (auto simp: inner_Basis)
moreover
- { fix x::"'a" assume "x\<in>?A"
- hence "finite d" "x\<in>?A" using assms by(auto intro:finite_subset)
- hence "x\<in> span ?B"
+ { fix x::"'a" assume "x \<in> ?A"
+ hence "finite d" "x \<in> ?A" using assms by(auto intro: finite_subset[OF _ finite_Basis])
+ from this d have "x \<in> span d"
proof(induct d arbitrary: x)
- case empty hence "x=0" apply(subst euclidean_eq) by auto
+ case empty hence "x=0" apply(rule_tac euclidean_eqI) by auto
thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
next
case (insert k F)
- hence *:"\<forall>i<DIM('a). i \<notin> insert k F \<longrightarrow> x $$ i = 0" by auto
+ hence *:"\<forall>i\<in>Basis. i \<notin> insert k F \<longrightarrow> x \<bullet> i = 0" by auto
have **:"F \<subseteq> insert k F" by auto
- def y \<equiv> "x - x$$k *\<^sub>R basis k"
- have y:"x = y + (x$$k) *\<^sub>R basis k" unfolding y_def by auto
- { fix i assume i':"i \<notin> F"
- hence "y $$ i = 0" unfolding y_def
- using *[THEN spec[where x=i]] by auto }
- hence "y \<in> span (basis ` F)" using insert(3) by auto
- hence "y \<in> span (basis ` (insert k F))"
- using span_mono[of "?bas ` F" "?bas ` (insert k F)"]
- using image_mono[OF **, of basis] using assms by auto
+ def y \<equiv> "x - (x\<bullet>k) *\<^sub>R k"
+ have y:"x = y + (x\<bullet>k) *\<^sub>R k" unfolding y_def by auto
+ { fix i assume i': "i \<notin> F" "i \<in> Basis"
+ hence "y \<bullet> i = 0" unfolding y_def
+ using *[THEN bspec[where x=i]] insert by (auto simp: inner_simps inner_Basis) }
+ hence "y \<in> span F" using insert by auto
+ hence "y \<in> span (insert k F)"
+ using span_mono[of F "insert k F"] using assms by auto
moreover
- have "basis k \<in> span (?bas ` (insert k F))" by(rule span_superset, auto)
- hence "x$$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
+ have "k \<in> span (insert k F)" by(rule span_superset, auto)
+ hence "(x\<bullet>k) *\<^sub>R k \<in> span (insert k F)"
using span_mul by auto
ultimately
- have "y + x$$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
+ have "y + (x\<bullet>k) *\<^sub>R k \<in> span (insert k F)"
using span_add by auto
thus ?case using y by auto
qed
}
- hence "?A \<subseteq> span ?B" by auto
+ hence "?A \<subseteq> span d" by auto
moreover
- { fix x assume "x \<in> ?B"
- hence "x\<in>{(basis i)::'a |i. i \<in> ?D}" using assms by auto }
- hence "independent ?B" using independent_mono[OF independent_basis, of ?B] and assms by auto
+ { fix x assume "x \<in> d" hence "x \<in> ?D" using assms by auto }
+ hence "independent d" using independent_mono[OF independent_Basis, of d] and assms by auto
moreover
have "d \<subseteq> ?D" unfolding subset_eq using assms by auto
- hence *:"inj_on (basis::nat\<Rightarrow>'a) d" using subset_inj_on[OF basis_inj, of "d"] by auto
- have "card ?B = card d" unfolding card_image[OF *] by auto
- ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto
+ ultimately show ?thesis using dim_unique[of d ?A] by auto
qed
text{* Hence closure and completeness of all subspaces. *}
-lemma closed_subspace_lemma: "n \<le> card (UNIV::'n::finite set) \<Longrightarrow> \<exists>A::'n set. card A = n"
-apply (induct n)
-apply (rule_tac x="{}" in exI, simp)
-apply clarsimp
-apply (subgoal_tac "\<exists>x. x \<notin> A")
-apply (erule exE)
-apply (rule_tac x="insert x A" in exI, simp)
-apply (subgoal_tac "A \<noteq> UNIV", auto)
-done
+lemma ex_card: assumes "n \<le> card A" shows "\<exists>S\<subseteq>A. card S = n"
+proof cases
+ assume "finite A"
+ from ex_bij_betw_nat_finite[OF this] guess f ..
+ moreover with `n \<le> card A` have "{..< n} \<subseteq> {..< card A}" "inj_on f {..< n}"
+ by (auto simp: bij_betw_def intro: subset_inj_on)
+ ultimately have "f ` {..< n} \<subseteq> A" "card (f ` {..< n}) = n"
+ by (auto simp: bij_betw_def card_image)
+ then show ?thesis by blast
+next
+ assume "\<not> finite A" with `n \<le> card A` show ?thesis by force
+qed
lemma closed_subspace: fixes s::"('a::euclidean_space) set"
assumes "subspace s" shows "closed s"
proof-
- have *:"dim s \<le> DIM('a)" using dim_subset_UNIV by auto
- def d \<equiv> "{..<dim s}" have t:"card d = dim s" unfolding d_def by auto
- let ?t = "{x::'a. \<forall>i<DIM('a). i \<notin> d \<longrightarrow> x$$i = 0}"
- have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i<DIM('a). i \<notin> d \<longrightarrow> x $$ i = 0} = s \<and>
- inj_on f {x::'a. \<forall>i<DIM('a). i \<notin> d \<longrightarrow> x $$ i = 0}"
- apply(rule subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]])
- using dim_substandard[of d,where 'a='a] and t unfolding d_def using * assms by auto
- then guess f apply-by(erule exE conjE)+ note f = this
+ have "dim s \<le> card (Basis :: 'a set)" using dim_subset_UNIV by auto
+ with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \<subseteq> Basis" by auto
+ let ?t = "{x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
+ have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s \<and>
+ inj_on f {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
+ using dim_substandard[of d] t d assms
+ by (intro subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]]) (auto simp: inner_Basis)
+ then guess f by (elim exE conjE) note f = this
interpret f: bounded_linear f using f unfolding linear_conv_bounded_linear by auto
- have "\<forall>x\<in>?t. f x = 0 \<longrightarrow> x = 0" using f.zero using f(3)[unfolded inj_on_def]
- by(erule_tac x=0 in ballE) auto
+ { fix x have "x\<in>?t \<Longrightarrow> f x = 0 \<Longrightarrow> x = 0" using f.zero d f(3)[THEN inj_onD, of x 0] by auto }
moreover have "closed ?t" using closed_substandard .
moreover have "subspace ?t" using subspace_substandard .
ultimately show ?thesis using closed_injective_image_subspace[of ?t f]
@@ -6144,7 +6090,7 @@
proof(cases "m=0")
{ fix x assume "x \<le> c" "c \<le> x"
hence "x=c" unfolding eucl_le[where 'a='a] apply-
- apply(subst euclidean_eq) by (auto intro: order_antisym) }
+ apply(subst euclidean_eq_iff) by (auto intro: order_antisym) }
moreover case True
moreover have "c \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: eucl_le[where 'a='a])
ultimately show ?thesis by auto
@@ -6152,23 +6098,23 @@
case False
{ fix y assume "a \<le> y" "y \<le> b" "m > 0"
hence "m *\<^sub>R a + c \<le> m *\<^sub>R y + c" "m *\<^sub>R y + c \<le> m *\<^sub>R b + c"
- unfolding eucl_le[where 'a='a] by auto
+ unfolding eucl_le[where 'a='a] by (auto simp: inner_simps)
} moreover
{ fix y assume "a \<le> y" "y \<le> b" "m < 0"
hence "m *\<^sub>R b + c \<le> m *\<^sub>R y + c" "m *\<^sub>R y + c \<le> m *\<^sub>R a + c"
- unfolding eucl_le[where 'a='a] by(auto simp add: mult_left_mono_neg)
+ unfolding eucl_le[where 'a='a] by(auto simp add: mult_left_mono_neg inner_simps)
} moreover
{ fix y assume "m > 0" "m *\<^sub>R a + c \<le> y" "y \<le> m *\<^sub>R b + c"
hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"])
- by(auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff)
+ by(auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_simps)
} moreover
{ fix y assume "m *\<^sub>R b + c \<le> y" "y \<le> m *\<^sub>R a + c" "m < 0"
hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"])
- by(auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff)
+ by(auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_simps)
}
ultimately show ?thesis using False by auto
qed
--- a/src/HOL/Probability/Borel_Space.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Probability/Borel_Space.thy Fri Dec 14 15:46:01 2012 +0100
@@ -126,344 +126,6 @@
apply(rule borel_measurableI)
using continuous_open_preimage[OF assms] unfolding vimage_def by auto
-section "Borel spaces on euclidean spaces"
-
-lemma borel_measurable_euclidean_component'[measurable]:
- "(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel"
- by (intro continuous_on_euclidean_component continuous_on_id borel_measurable_continuous_on1)
-
-lemma borel_measurable_euclidean_component:
- "(f :: 'a \<Rightarrow> 'b::euclidean_space) \<in> borel_measurable M \<Longrightarrow>(\<lambda>x. f x $$ i) \<in> borel_measurable M"
- by simp
-
-lemma [measurable]:
- fixes a b :: "'a\<Colon>ordered_euclidean_space"
- shows lessThan_borel: "{..< a} \<in> sets borel"
- and greaterThan_borel: "{a <..} \<in> sets borel"
- and greaterThanLessThan_borel: "{a<..<b} \<in> sets borel"
- and atMost_borel: "{..a} \<in> sets borel"
- and atLeast_borel: "{a..} \<in> sets borel"
- and atLeastAtMost_borel: "{a..b} \<in> sets borel"
- and greaterThanAtMost_borel: "{a<..b} \<in> sets borel"
- and atLeastLessThan_borel: "{a..<b} \<in> sets borel"
- unfolding greaterThanAtMost_def atLeastLessThan_def
- by (blast intro: borel_open borel_closed)+
-
-lemma
- shows hafspace_less_borel: "{x::'a::euclidean_space. a < x $$ i} \<in> sets borel"
- and hafspace_greater_borel: "{x::'a::euclidean_space. x $$ i < a} \<in> sets borel"
- and hafspace_less_eq_borel: "{x::'a::euclidean_space. a \<le> x $$ i} \<in> sets borel"
- and hafspace_greater_eq_borel: "{x::'a::euclidean_space. x $$ i \<le> a} \<in> sets borel"
- by simp_all
-
-lemma borel_measurable_less[measurable]:
- fixes f :: "'a \<Rightarrow> real"
- assumes f: "f \<in> borel_measurable M"
- assumes g: "g \<in> borel_measurable M"
- shows "{w \<in> space M. f w < g w} \<in> sets M"
-proof -
- have "{w \<in> space M. f w < g w} = {x \<in> space M. \<exists>r. f x < of_rat r \<and> of_rat r < g x}"
- using Rats_dense_in_real by (auto simp add: Rats_def)
- with f g show ?thesis
- by simp
-qed
-
-lemma
- fixes f :: "'a \<Rightarrow> real"
- assumes f[measurable]: "f \<in> borel_measurable M"
- assumes g[measurable]: "g \<in> borel_measurable M"
- shows borel_measurable_le[measurable]: "{w \<in> space M. f w \<le> g w} \<in> sets M"
- and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M"
- and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
- unfolding eq_iff not_less[symmetric]
- by measurable
-
-subsection "Borel space equals sigma algebras over intervals"
-
-lemma borel_sigma_sets_subset:
- "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"
- using sets.sigma_sets_subset[of A borel] by simp
-
-lemma borel_eq_sigmaI1:
- fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
- assumes borel_eq: "borel = sigma UNIV X"
- assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (range F))"
- assumes F: "\<And>i. F i \<in> sets borel"
- shows "borel = sigma UNIV (range F)"
- unfolding borel_def
-proof (intro sigma_eqI antisym)
- have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
- unfolding borel_def by simp
- also have "\<dots> = sigma_sets UNIV X"
- unfolding borel_eq by simp
- also have "\<dots> \<subseteq> sigma_sets UNIV (range F)"
- using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto
- finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (range F)" .
- show "sigma_sets UNIV (range F) \<subseteq> sigma_sets UNIV {S. open S}"
- unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
-qed auto
-
-lemma borel_eq_sigmaI2:
- fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set"
- and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
- assumes borel_eq: "borel = sigma UNIV (range (\<lambda>(i, j). G i j))"
- assumes X: "\<And>i j. G i j \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
- assumes F: "\<And>i j. F i j \<in> sets borel"
- shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
- using assms by (intro borel_eq_sigmaI1[where X="range (\<lambda>(i, j). G i j)" and F="(\<lambda>(i, j). F i j)"]) auto
-
-lemma borel_eq_sigmaI3:
- fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
- assumes borel_eq: "borel = sigma UNIV X"
- assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
- assumes F: "\<And>i j. F i j \<in> sets borel"
- shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
- using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto
-
-lemma borel_eq_sigmaI4:
- fixes F :: "'i \<Rightarrow> 'a::topological_space set"
- and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
- assumes borel_eq: "borel = sigma UNIV (range (\<lambda>(i, j). G i j))"
- assumes X: "\<And>i j. G i j \<in> sets (sigma UNIV (range F))"
- assumes F: "\<And>i. F i \<in> sets borel"
- shows "borel = sigma UNIV (range F)"
- using assms by (intro borel_eq_sigmaI1[where X="range (\<lambda>(i, j). G i j)" and F=F]) auto
-
-lemma borel_eq_sigmaI5:
- fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set"
- assumes borel_eq: "borel = sigma UNIV (range G)"
- assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
- assumes F: "\<And>i j. F i j \<in> sets borel"
- shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
- using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto
-
-lemma halfspace_gt_in_halfspace:
- "{x\<Colon>'a. a < x $$ i} \<in> sigma_sets UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))"
- (is "?set \<in> ?SIGMA")
-proof -
- interpret sigma_algebra UNIV ?SIGMA
- by (intro sigma_algebra_sigma_sets) simp_all
- have *: "?set = (\<Union>n. UNIV - {x\<Colon>'a. x $$ i < a + 1 / real (Suc n)})"
- proof (safe, simp_all add: not_less)
- fix x :: 'a assume "a < x $$ i"
- with reals_Archimedean[of "x $$ i - a"]
- obtain n where "a + 1 / real (Suc n) < x $$ i"
- by (auto simp: inverse_eq_divide field_simps)
- then show "\<exists>n. a + 1 / real (Suc n) \<le> x $$ i"
- by (blast intro: less_imp_le)
- next
- fix x n
- have "a < a + 1 / real (Suc n)" by auto
- also assume "\<dots> \<le> x"
- finally show "a < x" .
- qed
- show "?set \<in> ?SIGMA" unfolding *
- by (auto del: Diff intro!: Diff)
-qed
-
-lemma borel_eq_halfspace_less:
- "borel = sigma UNIV (range (\<lambda>(a, i). {x::'a::ordered_euclidean_space. x $$ i < a}))"
- (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI3[OF borel_def])
- fix S :: "'a set" assume "S \<in> {S. open S}"
- then have "open S" by simp
- from open_UNION[OF this]
- obtain I where *: "S =
- (\<Union>(a, b)\<in>I.
- (\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
- (\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
- unfolding greaterThanLessThan_def
- unfolding eucl_greaterThan_eq_halfspaces[where 'a='a]
- unfolding eucl_lessThan_eq_halfspaces[where 'a='a]
- by blast
- show "S \<in> ?SIGMA"
- unfolding *
- by (safe intro!: sets.countable_UN sets.Int sets.countable_INT)
- (auto intro!: halfspace_gt_in_halfspace)
-qed auto
-
-lemma borel_eq_halfspace_le:
- "borel = sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i \<le> a}))"
- (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
- fix a i
- have *: "{x::'a. x$$i < a} = (\<Union>n. {x. x$$i \<le> a - 1/real (Suc n)})"
- proof (safe, simp_all)
- fix x::'a assume *: "x$$i < a"
- with reals_Archimedean[of "a - x$$i"]
- obtain n where "x $$ i < a - 1 / (real (Suc n))"
- by (auto simp: field_simps inverse_eq_divide)
- then show "\<exists>n. x $$ i \<le> a - 1 / (real (Suc n))"
- by (blast intro: less_imp_le)
- next
- fix x::'a and n
- assume "x$$i \<le> a - 1 / real (Suc n)"
- also have "\<dots> < a" by auto
- finally show "x$$i < a" .
- qed
- show "{x. x$$i < a} \<in> ?SIGMA" unfolding *
- by (safe intro!: sets.countable_UN) auto
-qed auto
-
-lemma borel_eq_halfspace_ge:
- "borel = sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a \<le> x $$ i}))"
- (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
- fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
- show "{x. x$$i < a} \<in> ?SIGMA" unfolding *
- by (safe intro!: sets.compl_sets) auto
-qed auto
-
-lemma borel_eq_halfspace_greater:
- "borel = sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a < x $$ i}))"
- (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
- fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
- show "{x. x$$i \<le> a} \<in> ?SIGMA" unfolding *
- by (safe intro!: sets.compl_sets) auto
-qed auto
-
-lemma borel_eq_atMost:
- "borel = sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))"
- (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
- fix a i show "{x. x$$i \<le> a} \<in> ?SIGMA"
- proof cases
- assume "i < DIM('a)"
- then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
- proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm)
- fix x
- from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
- then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
- by (subst (asm) Max_le_iff) auto
- then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
- by (auto intro!: exI[of _ k])
- qed
- show "{x. x$$i \<le> a} \<in> ?SIGMA" unfolding *
- by (safe intro!: sets.countable_UN) auto
- qed (auto intro: sigma_sets_top sigma_sets.Empty)
-qed auto
-
-lemma borel_eq_greaterThan:
- "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {a<..}))"
- (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
- fix a i show "{x. x$$i \<le> a} \<in> ?SIGMA"
- proof cases
- assume "i < DIM('a)"
- have "{x::'a. x$$i \<le> a} = UNIV - {x::'a. a < x$$i}" by auto
- also have *: "{x::'a. a < x$$i} = (\<Union>k::nat. {(\<chi>\<chi> n. if n = i then a else -real k) <..})" using `i <DIM('a)`
- proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
- fix x
- from reals_Archimedean2[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
- guess k::nat .. note k = this
- { fix i assume "i < DIM('a)"
- then have "-x$$i < real k"
- using k by (subst (asm) Max_less_iff) auto
- then have "- real k < x$$i" by simp }
- then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
- by (auto intro!: exI[of _ k])
- qed
- finally show "{x. x$$i \<le> a} \<in> ?SIGMA"
- apply (simp only:)
- apply (safe intro!: sets.countable_UN sets.Diff)
- apply (auto intro: sigma_sets_top)
- done
- qed (auto intro: sigma_sets_top sigma_sets.Empty)
-qed auto
-
-lemma borel_eq_lessThan:
- "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {..<a}))"
- (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
- fix a i show "{x. a \<le> x$$i} \<in> ?SIGMA"
- proof cases
- fix a i assume "i < DIM('a)"
- have "{x::'a. a \<le> x$$i} = UNIV - {x::'a. x$$i < a}" by auto
- also have *: "{x::'a. x$$i < a} = (\<Union>k::nat. {..< (\<chi>\<chi> n. if n = i then a else real k)})" using `i <DIM('a)`
- proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
- fix x
- from reals_Archimedean2[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"]
- guess k::nat .. note k = this
- { fix i assume "i < DIM('a)"
- then have "x$$i < real k"
- using k by (subst (asm) Max_less_iff) auto
- then have "x$$i < real k" by simp }
- then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia < real k"
- by (auto intro!: exI[of _ k])
- qed
- finally show "{x. a \<le> x$$i} \<in> ?SIGMA"
- apply (simp only:)
- apply (safe intro!: sets.countable_UN sets.Diff)
- apply (auto intro: sigma_sets_top)
- done
- qed (auto intro: sigma_sets_top sigma_sets.Empty)
-qed auto
-
-lemma borel_eq_atLeastAtMost:
- "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} \<Colon>'a\<Colon>ordered_euclidean_space set))"
- (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
- fix a::'a
- have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
- proof (safe, simp_all add: eucl_le[where 'a='a])
- fix x
- from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
- guess k::nat .. note k = this
- { fix i assume "i < DIM('a)"
- with k have "- x$$i \<le> real k"
- by (subst (asm) Max_le_iff) (auto simp: field_simps)
- then have "- real k \<le> x$$i" by simp }
- then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
- by (auto intro!: exI[of _ k])
- qed
- show "{..a} \<in> ?SIGMA" unfolding *
- by (safe intro!: sets.countable_UN)
- (auto intro!: sigma_sets_top)
-qed auto
-
-lemma borel_eq_greaterThanLessThan:
- "borel = sigma UNIV (range (\<lambda> (a, b). {a <..< b} :: 'a \<Colon> ordered_euclidean_space set))"
- (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI1[OF borel_def])
- fix M :: "'a set" assume "M \<in> {S. open S}"
- then have "open M" by simp
- show "M \<in> ?SIGMA"
- apply (subst open_UNION[OF `open M`])
- apply (safe intro!: sets.countable_UN)
- apply auto
- done
-qed auto
-
-lemma borel_eq_atLeastLessThan:
- "borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan])
- have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto
- fix x :: real
- have "{..<x} = (\<Union>i::nat. {-real i ..< x})"
- by (auto simp: move_uminus real_arch_simple)
- then show "{..< x} \<in> ?SIGMA"
- by (auto intro: sigma_sets.intros)
-qed auto
-
-lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)"
- unfolding borel_def
-proof (intro sigma_eqI sigma_sets_eqI, safe)
- fix x :: "'a set" assume "open x"
- hence "x = UNIV - (UNIV - x)" by auto
- also have "\<dots> \<in> sigma_sets UNIV (Collect closed)"
- by (rule sigma_sets.Compl)
- (auto intro!: sigma_sets.Basic simp: `open x`)
- finally show "x \<in> sigma_sets UNIV (Collect closed)" by simp
-next
- fix x :: "'a set" assume "closed x"
- hence "x = UNIV - (UNIV - x)" by auto
- also have "\<dots> \<in> sigma_sets UNIV (Collect open)"
- by (rule sigma_sets.Compl)
- (auto intro!: sigma_sets.Basic simp: `closed x`)
- finally show "x \<in> sigma_sets UNIV (Collect open)" by simp
-qed simp_all
-
lemma borel_eq_countable_basis:
fixes B::"'a::topological_space set set"
assumes "countable B"
@@ -491,82 +153,62 @@
"borel = sigma UNIV union_closed_basis"
by (rule borel_eq_countable_basis[OF countable_union_closed_basis basis_union_closed_basis])
-lemma borel_measurable_halfspacesI:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
- assumes F: "borel = sigma UNIV (range F)"
- and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
- and S: "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
- shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
-proof safe
- fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
- then show "S a i \<in> sets M" unfolding assms
- by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1))
-next
- assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
- { fix a i have "S a i \<in> sets M"
- proof cases
- assume "i < DIM('c)"
- with a show ?thesis unfolding assms(2) by simp
- next
- assume "\<not> i < DIM('c)"
- from S[OF this] show ?thesis .
- qed }
- then show "f \<in> borel_measurable M"
- by (auto intro!: measurable_measure_of simp: S_eq F)
+lemma topological_basis_prod:
+ assumes A: "topological_basis A" and B: "topological_basis B"
+ shows "topological_basis ((\<lambda>(a, b). a \<times> b) ` (A \<times> B))"
+ unfolding topological_basis_def
+proof (safe, simp_all del: ex_simps add: subset_image_iff ex_simps(1)[symmetric])
+ fix S :: "('a \<times> 'b) set" assume "open S"
+ then show "\<exists>X\<subseteq>A \<times> B. (\<Union>(a,b)\<in>X. a \<times> b) = S"
+ proof (safe intro!: exI[of _ "{x\<in>A \<times> B. fst x \<times> snd x \<subseteq> S}"])
+ fix x y assume "(x, y) \<in> S"
+ from open_prod_elim[OF `open S` this]
+ obtain a b where a: "open a""x \<in> a" and b: "open b" "y \<in> b" and "a \<times> b \<subseteq> S"
+ by (metis mem_Sigma_iff)
+ moreover from topological_basisE[OF A a] guess A0 .
+ moreover from topological_basisE[OF B b] guess B0 .
+ ultimately show "(x, y) \<in> (\<Union>(a, b)\<in>{X \<in> A \<times> B. fst X \<times> snd X \<subseteq> S}. a \<times> b)"
+ by (intro UN_I[of "(A0, B0)"]) auto
+ qed auto
+qed (metis A B topological_basis_open open_Times)
+
+instance prod :: (countable_basis_space, countable_basis_space) countable_basis_space
+proof
+ obtain A :: "'a set set" where "countable A" "topological_basis A"
+ using ex_countable_basis by auto
+ moreover
+ obtain B :: "'b set set" where "countable B" "topological_basis B"
+ using ex_countable_basis by auto
+ ultimately show "\<exists>B::('a \<times> 'b) set set. countable B \<and> topological_basis B"
+ by (auto intro!: exI[of _ "(\<lambda>(a, b). a \<times> b) ` (A \<times> B)"] topological_basis_prod)
qed
-lemma borel_measurable_iff_halfspace_le:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
- shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i \<le> a} \<in> sets M)"
- by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto
-
-lemma borel_measurable_iff_halfspace_less:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
- shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i < a} \<in> sets M)"
- by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto
-
-lemma borel_measurable_iff_halfspace_ge:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
- shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a \<le> f w $$ i} \<in> sets M)"
- by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto
-
-lemma borel_measurable_iff_halfspace_greater:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
- shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a < f w $$ i} \<in> sets M)"
- by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto
-
-lemma borel_measurable_iff_le:
- "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
- using borel_measurable_iff_halfspace_le[where 'c=real] by simp
+lemma borel_measurable_Pair[measurable (raw)]:
+ fixes f :: "'a \<Rightarrow> 'b::countable_basis_space" and g :: "'a \<Rightarrow> 'c::countable_basis_space"
+ assumes f[measurable]: "f \<in> borel_measurable M"
+ assumes g[measurable]: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M"
+proof (subst borel_eq_countable_basis)
+ let ?B = "SOME B::'b set set. countable B \<and> topological_basis B"
+ let ?C = "SOME B::'c set set. countable B \<and> topological_basis B"
+ let ?P = "(\<lambda>(b, c). b \<times> c) ` (?B \<times> ?C)"
+ show "countable ?P" "topological_basis ?P"
+ by (auto intro!: countable_basis topological_basis_prod is_basis)
-lemma borel_measurable_iff_less:
- "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
- using borel_measurable_iff_halfspace_less[where 'c=real] by simp
-
-lemma borel_measurable_iff_ge:
- "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
- using borel_measurable_iff_halfspace_ge[where 'c=real]
- by simp
-
-lemma borel_measurable_iff_greater:
- "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
- using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
-
-lemma borel_measurable_euclidean_space:
- fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
- shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M)"
-proof safe
- fix i assume "f \<in> borel_measurable M"
- then show "(\<lambda>x. f x $$ i) \<in> borel_measurable M"
- by (auto intro: borel_measurable_euclidean_component)
-next
- assume f: "\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M"
- then show "f \<in> borel_measurable M"
- unfolding borel_measurable_iff_halfspace_le by auto
+ show "(\<lambda>x. (f x, g x)) \<in> measurable M (sigma UNIV ?P)"
+ proof (rule measurable_measure_of)
+ fix S assume "S \<in> ?P"
+ then obtain b c where "b \<in> ?B" "c \<in> ?C" and S: "S = b \<times> c" by auto
+ then have borel: "open b" "open c"
+ by (auto intro: is_basis topological_basis_open)
+ have "(\<lambda>x. (f x, g x)) -` S \<inter> space M = (f -` b \<inter> space M) \<inter> (g -` c \<inter> space M)"
+ unfolding S by auto
+ also have "\<dots> \<in> sets M"
+ using borel by simp
+ finally show "(\<lambda>x. (f x, g x)) -` S \<inter> space M \<in> sets M" .
+ qed auto
qed
-subsection "Borel measurable operators"
-
lemma borel_measurable_continuous_on:
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
assumes f: "continuous_on UNIV f" and g: "g \<in> borel_measurable M"
@@ -598,30 +240,6 @@
using measurable_comp[OF g borel_measurable_continuous_on_open'[OF cont], of c]
by (simp add: comp_def)
-lemma borel_measurable_uminus[measurable (raw)]:
- fixes g :: "'a \<Rightarrow> real"
- assumes g: "g \<in> borel_measurable M"
- shows "(\<lambda>x. - g x) \<in> borel_measurable M"
- by (rule borel_measurable_continuous_on[OF _ g]) (auto intro: continuous_on_minus continuous_on_id)
-
-lemma euclidean_component_prod:
- fixes x :: "'a :: euclidean_space \<times> 'b :: euclidean_space"
- shows "x $$ i = (if i < DIM('a) then fst x $$ i else snd x $$ (i - DIM('a)))"
- unfolding euclidean_component_def basis_prod_def inner_prod_def by auto
-
-lemma borel_measurable_Pair[measurable (raw)]:
- fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space" and g :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
- assumes f: "f \<in> borel_measurable M"
- assumes g: "g \<in> borel_measurable M"
- shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M"
-proof (intro borel_measurable_iff_halfspace_le[THEN iffD2] allI impI)
- fix i and a :: real assume i: "i < DIM('b \<times> 'c)"
- have [simp]: "\<And>P A B C. {w. (P \<longrightarrow> A w \<and> B w) \<and> (\<not> P \<longrightarrow> A w \<and> C w)} =
- {w. A w \<and> (P \<longrightarrow> B w) \<and> (\<not> P \<longrightarrow> C w)}" by auto
- from i f g show "{w \<in> space M. (f w, g w) $$ i \<le> a} \<in> sets M"
- by (auto simp: euclidean_component_prod)
-qed
-
lemma continuous_on_fst: "continuous_on UNIV fst"
proof -
have [simp]: "range fst = UNIV" by (auto simp: image_iff)
@@ -639,7 +257,7 @@
qed
lemma borel_measurable_continuous_Pair:
- fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space" and g :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
+ fixes f :: "'a \<Rightarrow> 'b::countable_basis_space" and g :: "'a \<Rightarrow> 'c::countable_basis_space"
assumes [measurable]: "f \<in> borel_measurable M"
assumes [measurable]: "g \<in> borel_measurable M"
assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))"
@@ -650,6 +268,413 @@
unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto
qed
+section "Borel spaces on euclidean spaces"
+
+lemma borel_measurable_inner[measurable (raw)]:
+ fixes f g :: "'a \<Rightarrow> 'b::{countable_basis_space, real_inner}"
+ assumes "f \<in> borel_measurable M"
+ assumes "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x \<bullet> g x) \<in> borel_measurable M"
+ using assms
+ by (rule borel_measurable_continuous_Pair)
+ (intro continuous_on_inner continuous_on_snd continuous_on_fst)
+
+lemma [measurable]:
+ fixes a b :: "'a\<Colon>ordered_euclidean_space"
+ shows lessThan_borel: "{..< a} \<in> sets borel"
+ and greaterThan_borel: "{a <..} \<in> sets borel"
+ and greaterThanLessThan_borel: "{a<..<b} \<in> sets borel"
+ and atMost_borel: "{..a} \<in> sets borel"
+ and atLeast_borel: "{a..} \<in> sets borel"
+ and atLeastAtMost_borel: "{a..b} \<in> sets borel"
+ and greaterThanAtMost_borel: "{a<..b} \<in> sets borel"
+ and atLeastLessThan_borel: "{a..<b} \<in> sets borel"
+ unfolding greaterThanAtMost_def atLeastLessThan_def
+ by (blast intro: borel_open borel_closed)+
+
+lemma borel_measurable_less[measurable]:
+ fixes f :: "'a \<Rightarrow> real"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "{w \<in> space M. f w < g w} \<in> sets M"
+proof -
+ have "{w \<in> space M. f w < g w} = {x \<in> space M. \<exists>r. f x < of_rat r \<and> of_rat r < g x}"
+ using Rats_dense_in_real by (auto simp add: Rats_def)
+ with f g show ?thesis
+ by simp
+qed
+
+lemma
+ fixes f :: "'a \<Rightarrow> real"
+ assumes f[measurable]: "f \<in> borel_measurable M"
+ assumes g[measurable]: "g \<in> borel_measurable M"
+ shows borel_measurable_le[measurable]: "{w \<in> space M. f w \<le> g w} \<in> sets M"
+ and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M"
+ and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
+ unfolding eq_iff not_less[symmetric]
+ by measurable
+
+lemma
+ shows hafspace_less_borel: "{x::'a::euclidean_space. a < x \<bullet> i} \<in> sets borel"
+ and hafspace_greater_borel: "{x::'a::euclidean_space. x \<bullet> i < a} \<in> sets borel"
+ and hafspace_less_eq_borel: "{x::'a::euclidean_space. a \<le> x \<bullet> i} \<in> sets borel"
+ and hafspace_greater_eq_borel: "{x::'a::euclidean_space. x \<bullet> i \<le> a} \<in> sets borel"
+ by simp_all
+
+subsection "Borel space equals sigma algebras over intervals"
+
+lemma borel_sigma_sets_subset:
+ "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"
+ using sets.sigma_sets_subset[of A borel] by simp
+
+lemma borel_eq_sigmaI1:
+ fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
+ assumes borel_eq: "borel = sigma UNIV X"
+ assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (F ` A))"
+ assumes F: "\<And>i. i \<in> A \<Longrightarrow> F i \<in> sets borel"
+ shows "borel = sigma UNIV (F ` A)"
+ unfolding borel_def
+proof (intro sigma_eqI antisym)
+ have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
+ unfolding borel_def by simp
+ also have "\<dots> = sigma_sets UNIV X"
+ unfolding borel_eq by simp
+ also have "\<dots> \<subseteq> sigma_sets UNIV (F`A)"
+ using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto
+ finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (F`A)" .
+ show "sigma_sets UNIV (F`A) \<subseteq> sigma_sets UNIV {S. open S}"
+ unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
+qed auto
+
+lemma borel_eq_sigmaI2:
+ fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set"
+ and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
+ assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`B)"
+ assumes X: "\<And>i j. (i, j) \<in> B \<Longrightarrow> G i j \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"
+ assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel"
+ shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"
+ using assms
+ by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` B" and F="(\<lambda>(i, j). F i j)"]) auto
+
+lemma borel_eq_sigmaI3:
+ fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
+ assumes borel_eq: "borel = sigma UNIV X"
+ assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"
+ assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel"
+ shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"
+ using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto
+
+lemma borel_eq_sigmaI4:
+ fixes F :: "'i \<Rightarrow> 'a::topological_space set"
+ and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
+ assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`A)"
+ assumes X: "\<And>i j. (i, j) \<in> A \<Longrightarrow> G i j \<in> sets (sigma UNIV (range F))"
+ assumes F: "\<And>i. F i \<in> sets borel"
+ shows "borel = sigma UNIV (range F)"
+ using assms by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` A" and F=F]) auto
+
+lemma borel_eq_sigmaI5:
+ fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set"
+ assumes borel_eq: "borel = sigma UNIV (range G)"
+ assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
+ assumes F: "\<And>i j. F i j \<in> sets borel"
+ shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
+ using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto
+
+lemma borel_eq_box:
+ "borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a \<Colon> euclidean_space set))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI1[OF borel_def])
+ fix M :: "'a set" assume "M \<in> {S. open S}"
+ then have "open M" by simp
+ show "M \<in> ?SIGMA"
+ apply (subst open_UNION_box[OF `open M`])
+ apply (safe intro!: sets.countable_UN' countable_PiE countable_Collect)
+ apply (auto intro: countable_rat)
+ done
+qed (auto simp: box_def)
+
+lemma borel_eq_greaterThanLessThan:
+ "borel = sigma UNIV (range (\<lambda> (a, b). {a <..< b} :: 'a \<Colon> ordered_euclidean_space set))"
+ unfolding borel_eq_box apply (rule arg_cong2[where f=sigma])
+ by (auto simp: box_def image_iff mem_interval set_eq_iff simp del: greaterThanLessThan_iff)
+
+lemma halfspace_gt_in_halfspace:
+ assumes i: "i \<in> A"
+ shows "{x\<Colon>'a. a < x \<bullet> i} \<in>
+ sigma_sets UNIV ((\<lambda> (a, i). {x\<Colon>'a\<Colon>euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> A))"
+ (is "?set \<in> ?SIGMA")
+proof -
+ interpret sigma_algebra UNIV ?SIGMA
+ by (intro sigma_algebra_sigma_sets) simp_all
+ have *: "?set = (\<Union>n. UNIV - {x\<Colon>'a. x \<bullet> i < a + 1 / real (Suc n)})"
+ proof (safe, simp_all add: not_less)
+ fix x :: 'a assume "a < x \<bullet> i"
+ with reals_Archimedean[of "x \<bullet> i - a"]
+ obtain n where "a + 1 / real (Suc n) < x \<bullet> i"
+ by (auto simp: inverse_eq_divide field_simps)
+ then show "\<exists>n. a + 1 / real (Suc n) \<le> x \<bullet> i"
+ by (blast intro: less_imp_le)
+ next
+ fix x n
+ have "a < a + 1 / real (Suc n)" by auto
+ also assume "\<dots> \<le> x"
+ finally show "a < x" .
+ qed
+ show "?set \<in> ?SIGMA" unfolding *
+ by (auto del: Diff intro!: Diff i)
+qed
+
+lemma borel_eq_halfspace_less:
+ "borel = sigma UNIV ((\<lambda>(a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> Basis))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_box])
+ fix a b :: 'a
+ have "box a b = {x\<in>space ?SIGMA. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}"
+ by (auto simp: box_def)
+ also have "\<dots> \<in> sets ?SIGMA"
+ by (intro sets.sets_Collect_conj sets.sets_Collect_finite_All sets.sets_Collect_const)
+ (auto intro!: halfspace_gt_in_halfspace countable_PiE countable_rat)
+ finally show "box a b \<in> sets ?SIGMA" .
+qed auto
+
+lemma borel_eq_halfspace_le:
+ "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i \<le> a}) ` (UNIV \<times> Basis))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
+ fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
+ then have i: "i \<in> Basis" by auto
+ have *: "{x::'a. x\<bullet>i < a} = (\<Union>n. {x. x\<bullet>i \<le> a - 1/real (Suc n)})"
+ proof (safe, simp_all)
+ fix x::'a assume *: "x\<bullet>i < a"
+ with reals_Archimedean[of "a - x\<bullet>i"]
+ obtain n where "x \<bullet> i < a - 1 / (real (Suc n))"
+ by (auto simp: field_simps inverse_eq_divide)
+ then show "\<exists>n. x \<bullet> i \<le> a - 1 / (real (Suc n))"
+ by (blast intro: less_imp_le)
+ next
+ fix x::'a and n
+ assume "x\<bullet>i \<le> a - 1 / real (Suc n)"
+ also have "\<dots> < a" by auto
+ finally show "x\<bullet>i < a" .
+ qed
+ show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
+ by (safe intro!: sets.countable_UN) (auto intro: i)
+qed auto
+
+lemma borel_eq_halfspace_ge:
+ "borel = sigma UNIV ((\<lambda> (a, i). {x\<Colon>'a\<Colon>euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
+ fix a :: real and i :: 'a assume i: "(a, i) \<in> UNIV \<times> Basis"
+ have *: "{x::'a. x\<bullet>i < a} = space ?SIGMA - {x::'a. a \<le> x\<bullet>i}" by auto
+ show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
+ using i by (safe intro!: sets.compl_sets) auto
+qed auto
+
+lemma borel_eq_halfspace_greater:
+ "borel = sigma UNIV ((\<lambda> (a, i). {x\<Colon>'a\<Colon>euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
+ fix a :: real and i :: 'a assume "(a, i) \<in> (UNIV \<times> Basis)"
+ then have i: "i \<in> Basis" by auto
+ have *: "{x::'a. x\<bullet>i \<le> a} = space ?SIGMA - {x::'a. a < x\<bullet>i}" by auto
+ show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
+ by (safe intro!: sets.compl_sets) (auto intro: i)
+qed auto
+
+lemma borel_eq_atMost:
+ "borel = sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
+ fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
+ then have "i \<in> Basis" by auto
+ then have *: "{x::'a. x\<bullet>i \<le> a} = (\<Union>k::nat. {.. (\<Sum>n\<in>Basis. (if n = i then a else real k)*\<^sub>R n)})"
+ proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm)
+ fix x :: 'a
+ from real_arch_simple[of "Max ((\<lambda>i. x\<bullet>i)`Basis)"] guess k::nat ..
+ then have "\<And>i. i \<in> Basis \<Longrightarrow> x\<bullet>i \<le> real k"
+ by (subst (asm) Max_le_iff) auto
+ then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> x \<bullet> ia \<le> real k"
+ by (auto intro!: exI[of _ k])
+ qed
+ show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
+ by (safe intro!: sets.countable_UN) auto
+qed auto
+
+lemma borel_eq_greaterThan:
+ "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {a<..}))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
+ fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
+ then have i: "i \<in> Basis" by auto
+ have "{x::'a. x\<bullet>i \<le> a} = UNIV - {x::'a. a < x\<bullet>i}" by auto
+ also have *: "{x::'a. a < x\<bullet>i} =
+ (\<Union>k::nat. {\<Sum>n\<in>Basis. (if n = i then a else -real k) *\<^sub>R n <..})" using i
+ proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
+ fix x :: 'a
+ from reals_Archimedean2[of "Max ((\<lambda>i. -x\<bullet>i)`Basis)"]
+ guess k::nat .. note k = this
+ { fix i :: 'a assume "i \<in> Basis"
+ then have "-x\<bullet>i < real k"
+ using k by (subst (asm) Max_less_iff) auto
+ then have "- real k < x\<bullet>i" by simp }
+ then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> -real k < x \<bullet> ia"
+ by (auto intro!: exI[of _ k])
+ qed
+ finally show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA"
+ apply (simp only:)
+ apply (safe intro!: sets.countable_UN sets.Diff)
+ apply (auto intro: sigma_sets_top)
+ done
+qed auto
+
+lemma borel_eq_lessThan:
+ "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {..<a}))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
+ fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
+ then have i: "i \<in> Basis" by auto
+ have "{x::'a. a \<le> x\<bullet>i} = UNIV - {x::'a. x\<bullet>i < a}" by auto
+ also have *: "{x::'a. x\<bullet>i < a} = (\<Union>k::nat. {..< \<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n})" using `i\<in> Basis`
+ proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
+ fix x :: 'a
+ from reals_Archimedean2[of "Max ((\<lambda>i. x\<bullet>i)`Basis)"]
+ guess k::nat .. note k = this
+ { fix i :: 'a assume "i \<in> Basis"
+ then have "x\<bullet>i < real k"
+ using k by (subst (asm) Max_less_iff) auto
+ then have "x\<bullet>i < real k" by simp }
+ then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> x \<bullet> ia < real k"
+ by (auto intro!: exI[of _ k])
+ qed
+ finally show "{x. a \<le> x\<bullet>i} \<in> ?SIGMA"
+ apply (simp only:)
+ apply (safe intro!: sets.countable_UN sets.Diff)
+ apply (auto intro: sigma_sets_top)
+ done
+qed auto
+
+lemma borel_eq_atLeastAtMost:
+ "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} \<Colon>'a\<Colon>ordered_euclidean_space set))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
+ fix a::'a
+ have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
+ proof (safe, simp_all add: eucl_le[where 'a='a])
+ fix x :: 'a
+ from real_arch_simple[of "Max ((\<lambda>i. - x\<bullet>i)`Basis)"]
+ guess k::nat .. note k = this
+ { fix i :: 'a assume "i \<in> Basis"
+ with k have "- x\<bullet>i \<le> real k"
+ by (subst (asm) Max_le_iff) (auto simp: field_simps)
+ then have "- real k \<le> x\<bullet>i" by simp }
+ then show "\<exists>n::nat. \<forall>i\<in>Basis. - real n \<le> x \<bullet> i"
+ by (auto intro!: exI[of _ k])
+ qed
+ show "{..a} \<in> ?SIGMA" unfolding *
+ by (safe intro!: sets.countable_UN)
+ (auto intro!: sigma_sets_top)
+qed auto
+
+lemma borel_eq_atLeastLessThan:
+ "borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan])
+ have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto
+ fix x :: real
+ have "{..<x} = (\<Union>i::nat. {-real i ..< x})"
+ by (auto simp: move_uminus real_arch_simple)
+ then show "{..< x} \<in> ?SIGMA"
+ by (auto intro: sigma_sets.intros)
+qed auto
+
+lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)"
+ unfolding borel_def
+proof (intro sigma_eqI sigma_sets_eqI, safe)
+ fix x :: "'a set" assume "open x"
+ hence "x = UNIV - (UNIV - x)" by auto
+ also have "\<dots> \<in> sigma_sets UNIV (Collect closed)"
+ by (rule sigma_sets.Compl)
+ (auto intro!: sigma_sets.Basic simp: `open x`)
+ finally show "x \<in> sigma_sets UNIV (Collect closed)" by simp
+next
+ fix x :: "'a set" assume "closed x"
+ hence "x = UNIV - (UNIV - x)" by auto
+ also have "\<dots> \<in> sigma_sets UNIV (Collect open)"
+ by (rule sigma_sets.Compl)
+ (auto intro!: sigma_sets.Basic simp: `closed x`)
+ finally show "x \<in> sigma_sets UNIV (Collect open)" by simp
+qed simp_all
+
+lemma borel_measurable_halfspacesI:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ assumes F: "borel = sigma UNIV (F ` (UNIV \<times> Basis))"
+ and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
+ shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a::real. S a i \<in> sets M)"
+proof safe
+ fix a :: real and i :: 'b assume i: "i \<in> Basis" and f: "f \<in> borel_measurable M"
+ then show "S a i \<in> sets M" unfolding assms
+ by (auto intro!: measurable_sets simp: assms(1))
+next
+ assume a: "\<forall>i\<in>Basis. \<forall>a. S a i \<in> sets M"
+ then show "f \<in> borel_measurable M"
+ by (auto intro!: measurable_measure_of simp: S_eq F)
+qed
+
+lemma borel_measurable_iff_halfspace_le:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i \<le> a} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto
+
+lemma borel_measurable_iff_halfspace_less:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i < a} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto
+
+lemma borel_measurable_iff_halfspace_ge:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a \<le> f w \<bullet> i} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto
+
+lemma borel_measurable_iff_halfspace_greater:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a < f w \<bullet> i} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto
+
+lemma borel_measurable_iff_le:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
+ using borel_measurable_iff_halfspace_le[where 'c=real] by simp
+
+lemma borel_measurable_iff_less:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
+ using borel_measurable_iff_halfspace_less[where 'c=real] by simp
+
+lemma borel_measurable_iff_ge:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
+ using borel_measurable_iff_halfspace_ge[where 'c=real]
+ by simp
+
+lemma borel_measurable_iff_greater:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
+ using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
+
+lemma borel_measurable_euclidean_space:
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
+ shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M)"
+proof safe
+ assume f: "\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M"
+ then show "f \<in> borel_measurable M"
+ by (subst borel_measurable_iff_halfspace_le) auto
+qed auto
+
+subsection "Borel measurable operators"
+
+lemma borel_measurable_uminus[measurable (raw)]:
+ fixes g :: "'a \<Rightarrow> real"
+ assumes g: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. - g x) \<in> borel_measurable M"
+ by (rule borel_measurable_continuous_on[OF _ g]) (auto intro: continuous_on_minus continuous_on_id)
+
lemma borel_measurable_add[measurable (raw)]:
fixes f g :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
assumes f: "f \<in> borel_measurable M"
@@ -772,7 +797,7 @@
lemma borel_measurable_nth[measurable (raw)]:
"(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel"
- by (simp add: nth_conv_component)
+ by (simp add: cart_eq_inner_axis)
lemma convex_measurable:
fixes a b :: real
--- a/src/HOL/Probability/Lebesgue_Measure.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Probability/Lebesgue_Measure.thy Fri Dec 14 15:46:01 2012 +0100
@@ -36,7 +36,7 @@
subsection {* Standard Cubes *}
definition cube :: "nat \<Rightarrow> 'a::ordered_euclidean_space set" where
- "cube n \<equiv> {\<chi>\<chi> i. - real n .. \<chi>\<chi> i. real n}"
+ "cube n \<equiv> {\<Sum>i\<in>Basis. - n *\<^sub>R i .. \<Sum>i\<in>Basis. n *\<^sub>R i}"
lemma borel_cube[intro]: "cube n \<in> sets borel"
unfolding cube_def by auto
@@ -45,53 +45,44 @@
unfolding cube_def by auto
lemma cube_subset[intro]: "n \<le> N \<Longrightarrow> cube n \<subseteq> cube N"
- by (fastforce simp: eucl_le[where 'a='a] cube_def)
+ by (fastforce simp: eucl_le[where 'a='a] cube_def setsum_negf)
+
+lemma cube_subset_iff: "cube n \<subseteq> cube N \<longleftrightarrow> n \<le> N"
+ unfolding cube_def subset_interval by (simp add: setsum_negf ex_in_conv)
-lemma cube_subset_iff:
- "cube n \<subseteq> cube N \<longleftrightarrow> n \<le> N"
-proof
- assume subset: "cube n \<subseteq> (cube N::'a set)"
- then have "((\<chi>\<chi> i. real n)::'a) \<in> cube N"
- using DIM_positive[where 'a='a]
- by (fastforce simp: cube_def eucl_le[where 'a='a])
- then show "n \<le> N"
- by (fastforce simp: cube_def eucl_le[where 'a='a])
-next
- assume "n \<le> N" then show "cube n \<subseteq> (cube N::'a set)" by (rule cube_subset)
-qed
-
-lemma ball_subset_cube:"ball (0::'a::ordered_euclidean_space) (real n) \<subseteq> cube n"
- unfolding cube_def subset_eq mem_interval apply safe unfolding euclidean_lambda_beta'
-proof- fix x::'a and i assume as:"x\<in>ball 0 (real n)" "i<DIM('a)"
- thus "- real n \<le> x $$ i" "real n \<ge> x $$ i"
- using component_le_norm[of x i] by(auto simp: dist_norm)
+lemma ball_subset_cube: "ball (0::'a::ordered_euclidean_space) (real n) \<subseteq> cube n"
+ apply (simp add: cube_def subset_eq mem_interval setsum_negf eucl_le[where 'a='a])
+proof safe
+ fix x i :: 'a assume x: "x \<in> ball 0 (real n)" and i: "i \<in> Basis"
+ thus "- real n \<le> x \<bullet> i" "real n \<ge> x \<bullet> i"
+ using Basis_le_norm[OF i, of x] by(auto simp: dist_norm)
qed
lemma mem_big_cube: obtains n where "x \<in> cube n"
-proof- from reals_Archimedean2[of "norm x"] guess n ..
- thus ?thesis apply-apply(rule that[where n=n])
- apply(rule ball_subset_cube[unfolded subset_eq,rule_format])
- by (auto simp add:dist_norm)
+proof -
+ from reals_Archimedean2[of "norm x"] guess n ..
+ with ball_subset_cube[unfolded subset_eq, of n]
+ show ?thesis
+ by (intro that[where n=n]) (auto simp add: dist_norm)
qed
lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)"
- unfolding cube_def subset_eq apply safe unfolding mem_interval apply auto done
+ unfolding cube_def subset_interval by (simp add: setsum_negf)
lemma has_integral_interval_cube:
fixes a b :: "'a::ordered_euclidean_space"
- shows "(indicator {a .. b} has_integral
- content ({\<chi>\<chi> i. max (- real n) (a $$ i) .. \<chi>\<chi> i. min (real n) (b $$ i)} :: 'a set)) (cube n)"
+ shows "(indicator {a .. b} has_integral content ({a .. b} \<inter> cube n)) (cube n)"
(is "(?I has_integral content ?R) (cube n)")
proof -
- let "{?N .. ?P}" = ?R
have [simp]: "(\<lambda>x. if x \<in> cube n then ?I x else 0) = indicator ?R"
by (auto simp: indicator_def cube_def fun_eq_iff eucl_le[where 'a='a])
have "(?I has_integral content ?R) (cube n) \<longleftrightarrow> (indicator ?R has_integral content ?R) UNIV"
unfolding has_integral_restrict_univ[where s="cube n", symmetric] by simp
- also have "\<dots> \<longleftrightarrow> ((\<lambda>x. 1) has_integral content ?R) ?R"
- unfolding indicator_def [abs_def] has_integral_restrict_univ ..
- finally show ?thesis
- using has_integral_const[of "1::real" "?N" "?P"] by simp
+ also have "\<dots> \<longleftrightarrow> ((\<lambda>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 "((\<lambda>x. 1) has_integral content ?R *\<^sub>R 1) ?R"
+ unfolding cube_def inter_interval by (rule has_integral_const)
+ finally show ?thesis .
qed
subsection {* Lebesgue measure *}
@@ -219,11 +210,8 @@
also have "\<dots> \<subseteq> sets lebesgue"
proof (safe intro!: sets.sigma_sets_subset lebesgueI)
fix n :: nat and a b :: 'a
- let ?N = "\<chi>\<chi> i. max (- real n) (a $$ i)"
- let ?P = "\<chi>\<chi> i. min (real n) (b $$ i)"
show "(indicator {a..b} :: 'a\<Rightarrow>real) integrable_on cube n"
- unfolding integrable_on_def
- using has_integral_interval_cube[of a b] by auto
+ unfolding integrable_on_def using has_integral_interval_cube[of a b] by auto
qed
finally show ?thesis .
qed
@@ -418,12 +406,13 @@
case (Suc n')
have "real n \<le> (2 * real n)^1" by auto
also have "(2 * real n)^1 \<le> (2 * real n) ^ DIM('a)"
- using Suc DIM_positive[where 'a='a] by (intro power_increasing) (auto simp: real_of_nat_Suc)
+ using Suc DIM_positive[where 'a='a]
+ by (intro power_increasing) (auto simp: real_of_nat_Suc simp del: DIM_positive)
finally show ?thesis .
qed }
ultimately show "ereal (real n) \<le> ereal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
using integral_const DIM_positive[where 'a='a]
- by (auto simp: cube_def content_closed_interval_cases setprod_constant)
+ by (auto simp: cube_def content_closed_interval_cases setprod_constant setsum_negf)
qed simp
lemma lmeasure_complete: "A \<subseteq> B \<Longrightarrow> B \<in> null_sets lebesgue \<Longrightarrow> A \<in> null_sets lebesgue"
@@ -901,24 +890,24 @@
subsection {* Equivalence between product spaces and euclidean spaces *}
-definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> (nat \<Rightarrow> real)" where
- "e2p x = (\<lambda>i\<in>{..<DIM('a)}. x$$i)"
+definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> ('a \<Rightarrow> real)" where
+ "e2p x = (\<lambda>i\<in>Basis. x \<bullet> i)"
-definition p2e :: "(nat \<Rightarrow> real) \<Rightarrow> 'a::ordered_euclidean_space" where
- "p2e x = (\<chi>\<chi> i. x i)"
+definition p2e :: "('a \<Rightarrow> real) \<Rightarrow> 'a::ordered_euclidean_space" where
+ "p2e x = (\<Sum>i\<in>Basis. x i *\<^sub>R i)"
lemma e2p_p2e[simp]:
- "x \<in> extensional {..<DIM('a)} \<Longrightarrow> e2p (p2e x::'a::ordered_euclidean_space) = x"
+ "x \<in> extensional Basis \<Longrightarrow> e2p (p2e x::'a::ordered_euclidean_space) = x"
by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def)
lemma p2e_e2p[simp]:
"p2e (e2p x) = (x::'a::ordered_euclidean_space)"
- by (auto simp: euclidean_eq[where 'a='a] p2e_def e2p_def)
+ by (auto simp: euclidean_eq_iff[where 'a='a] p2e_def e2p_def)
interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure"
by default
-interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "{..<n}" for n :: nat
+interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "Basis"
by default auto
lemma sets_product_borel:
@@ -930,13 +919,13 @@
qed (auto simp: borel_eq_lessThan reals_Archimedean2)
lemma measurable_e2p[measurable]:
- "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))"
+ "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^isub>M (i::'a)\<in>Basis. (lborel :: real measure))"
proof (rule measurable_sigma_sets[OF sets_product_borel])
- fix A :: "(nat \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i} |x. True} "
- then obtain x where "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i})" by auto
- then have "e2p -` A = {..< (\<chi>\<chi> i. x i) :: 'a}"
+ fix A :: "('a \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^isub>E (i::'a)\<in>Basis. {..<x i} |x. True} "
+ then obtain x where "A = (\<Pi>\<^isub>E (i::'a)\<in>Basis. {..<x i})" by auto
+ then have "e2p -` A = {..< (\<Sum>i\<in>Basis. x i *\<^sub>R i) :: 'a}"
using DIM_positive by (auto simp add: set_eq_iff e2p_def
- euclidean_eq[where 'a='a] eucl_less[where 'a='a])
+ euclidean_eq_iff[where 'a='a] eucl_less[where 'a='a])
then show "e2p -` A \<inter> space (borel::'a measure) \<in> sets borel" by simp
qed (auto simp: e2p_def)
@@ -945,34 +934,34 @@
lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel" by simp
lemma measurable_p2e[measurable]:
- "p2e \<in> measurable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))
+ "p2e \<in> measurable (\<Pi>\<^isub>M (i::'a)\<in>Basis. (lborel :: real measure))
(borel :: 'a::ordered_euclidean_space measure)"
(is "p2e \<in> measurable ?P _")
proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2])
- fix x i
- let ?A = "{w \<in> space ?P. (p2e w :: 'a) $$ i \<le> x}"
- assume "i < DIM('a)"
- then have "?A = (\<Pi>\<^isub>E j\<in>{..<DIM('a)}. if i = j then {.. x} else UNIV)"
+ fix x and i :: 'a
+ let ?A = "{w \<in> space ?P. (p2e w :: 'a) \<bullet> i \<le> x}"
+ assume "i \<in> Basis"
+ then have "?A = (\<Pi>\<^isub>E j\<in>Basis. if i = j then {.. x} else UNIV)"
using DIM_positive by (auto simp: space_PiM p2e_def PiE_def split: split_if_asm)
then show "?A \<in> sets ?P"
by auto
qed
lemma lborel_eq_lborel_space:
- "(lborel :: 'a measure) = distr (\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel) borel p2e"
+ "(lborel :: 'a measure) = distr (\<Pi>\<^isub>M (i::'a::ordered_euclidean_space)\<in>Basis. lborel) borel p2e"
(is "?B = ?D")
proof (rule lborel_eqI)
show "sets ?D = sets borel" by simp
- let ?P = "(\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel)"
+ let ?P = "(\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel)"
fix a b :: 'a
- have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {a $$ i .. b $$ i})"
+ have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^isub>E i\<in>Basis. {a \<bullet> i .. b \<bullet> i})"
by (auto simp: eucl_le[where 'a='a] p2e_def space_PiM PiE_def Pi_iff)
have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}"
proof cases
assume "{a..b} \<noteq> {}"
then have "a \<le> b"
by (simp add: interval_ne_empty eucl_le[where 'a='a])
- then have "emeasure lborel {a..b} = (\<Prod>x<DIM('a). emeasure lborel {a $$ x .. b $$ x})"
+ then have "emeasure lborel {a..b} = (\<Prod>x\<in>Basis. emeasure lborel {a \<bullet> x .. b \<bullet> x})"
by (auto simp: content_closed_interval eucl_le[where 'a='a]
intro!: setprod_ereal[symmetric])
also have "\<dots> = emeasure ?P (p2e -` {a..b} \<inter> space ?P)"
@@ -986,13 +975,12 @@
lemma borel_fubini_positiv_integral:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
assumes f: "f \<in> borel_measurable borel"
- shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel)"
+ shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel)"
by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f)
lemma borel_fubini_integrable:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
- shows "integrable lborel f \<longleftrightarrow>
- integrable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel) (\<lambda>x. f (p2e x))"
+ shows "integrable lborel f \<longleftrightarrow> integrable (\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel) (\<lambda>x. f (p2e x))"
(is "_ \<longleftrightarrow> integrable ?B ?f")
proof
assume "integrable lborel f"
@@ -1017,7 +1005,7 @@
lemma borel_fubini:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
assumes f: "f \<in> borel_measurable borel"
- shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel))"
+ shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel))"
using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def)
lemma integrable_on_borel_integrable:
@@ -1066,7 +1054,7 @@
using borel_measurable_continuous_on_open[of "{a <..< b }" f "\<lambda>x. x" borel 0]
by (auto intro!: measurable_If[where P="\<lambda>x. x = a"] measurable_If[where P="\<lambda>x. x = b"])
also have "?g = ?f"
- using `a \<le> b` by (auto intro!: ext split: split_indicator)
+ using `a \<le> b` by (intro ext) (auto split: split_indicator)
finally show "?f \<in> borel_measurable lborel"
by simp
qed
--- a/src/HOL/Probability/Sigma_Algebra.thy Fri Dec 14 14:46:01 2012 +0100
+++ b/src/HOL/Probability/Sigma_Algebra.thy Fri Dec 14 15:46:01 2012 +0100
@@ -322,6 +322,20 @@
ultimately show ?thesis by simp
qed
+lemma (in sigma_algebra) countable_UN':
+ fixes A :: "'i \<Rightarrow> 'a set"
+ assumes X: "countable X"
+ assumes A: "A`X \<subseteq> M"
+ shows "(\<Union>x\<in>X. A x) \<in> M"
+proof -
+ have "(\<Union>x\<in>X. A x) = (\<Union>i\<in>to_nat_on X ` X. A (from_nat_into X i))"
+ using X by auto
+ also have "\<dots> \<in> M"
+ using A X
+ by (intro countable_UN) auto
+ finally show ?thesis .
+qed
+
lemma (in sigma_algebra) countable_INT [intro]:
fixes A :: "'i::countable \<Rightarrow> 'a set"
assumes A: "A`X \<subseteq> M" "X \<noteq> {}"
@@ -335,6 +349,20 @@
ultimately show ?thesis by metis
qed
+lemma (in sigma_algebra) countable_INT':
+ fixes A :: "'i \<Rightarrow> 'a set"
+ assumes X: "countable X" "X \<noteq> {}"
+ assumes A: "A`X \<subseteq> M"
+ shows "(\<Inter>x\<in>X. A x) \<in> M"
+proof -
+ have "(\<Inter>x\<in>X. A x) = (\<Inter>i\<in>to_nat_on X ` X. A (from_nat_into X i))"
+ using X by auto
+ also have "\<dots> \<in> M"
+ using A X
+ by (intro countable_INT) auto
+ finally show ?thesis .
+qed
+
lemma ring_of_sets_Pow: "ring_of_sets sp (Pow sp)"
by (auto simp: ring_of_sets_iff)
@@ -365,6 +393,16 @@
with assms show ?thesis by auto
qed
+lemma (in sigma_algebra) sets_Collect_countable_Ex':
+ assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
+ assumes "countable I"
+ shows "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} \<in> M"
+proof -
+ have "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} = (\<Union>i\<in>I. {x\<in>\<Omega>. P i x})" by auto
+ with assms show ?thesis
+ by (auto intro!: countable_UN')
+qed
+
lemmas (in sigma_algebra) sets_Collect =
sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const
sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All