diff -r 5001ed24e129 -r d5d342611edb src/HOL/Multivariate_Analysis/Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Aug 23 17:46:13 2010 +0200 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Aug 23 19:35:57 2010 +0200 @@ -1934,6 +1934,16 @@ assumes eucl_le: "x \ y \ (\i < DIM('a). x $$ i \ y $$ i)" and eucl_less: "x < y \ (\i < DIM('a). x $$ i < y $$ i)" +lemma eucl_less_not_refl[simp, intro!]: "\ x < (x::'a::ordered_euclidean_space)" + unfolding eucl_less[where 'a='a] by auto + +lemma euclidean_trans[trans]: + fixes x y z :: "'a::ordered_euclidean_space" + shows "x < y \ y < z \ x < z" + and "x \ y \ y < z \ x < z" + and "x \ y \ y \ z \ x \ z" + by (force simp: eucl_less[where 'a='a] eucl_le[where 'a='a])+ + subsection {* Linearity and Bilinearity continued *} lemma linear_bounded: @@ -3388,4 +3398,150 @@ instance complex :: euclidean_space proof qed (auto simp add: basis_complex_def inner_complex_def) +section {* Products Spaces *} + +instantiation prod :: (real_basis, real_basis) real_basis +begin + +definition "basis i = (if i < DIM('a) then (basis i, 0) else (0, basis (i - DIM('a))))" + +instance +proof + let ?b = "basis :: nat \ 'a \ 'b" + let ?b_a = "basis :: nat \ 'a" + let ?b_b = "basis :: nat \ 'b" + + note image_range = + image_add_atLeastLessThan[symmetric, of 0 "DIM('a)" "DIM('b)", simplified] + + have split_range: + "{.. {DIM('a).. (?b_b ` {.. {0}" + unfolding image_range image_image basis_prod_def_raw range_basis + by (auto simp: zero_prod_def basis_eq_0_iff) + hence b_split: + "?b ` {.. {0} \ {0} \ (?b_b ` {.. {DIM('b)+DIM('a)..}" + by auto + + have range_b: "range ?b = ?prod \ {0}" + by (subst split_UNIV) (simp add: image_Un b_split b_0) + + have prod: "\f A B. setsum f (A \ B) = (\a\A. \b\B. f (a, b))" + by (simp add: setsum_cartesian_product) + + show "span (range ?b) = UNIV" + unfolding span_explicit range_b + proof safe + fix a::'a and b::'b + from in_span_basis[of a] in_span_basis[of b] + obtain Sa ua Sb ub where span: + "finite Sa" "Sa \ basis ` {..v\Sa. ua v *\<^sub>R v)" + "finite Sb" "Sb \ basis ` {..v\Sb. ub v *\<^sub>R v)" + unfolding span_explicit by auto + + let ?S = "((Sa - {0}) \ {0} \ {0} \ (Sb - {0}))" + have *: + "?S \ {v. fst v = 0} \ {v. snd v = 0} = {}" + "?S \ - {v. fst v = 0} \ {v. snd v = 0} = (Sa - {0}) \ {0}" + "?S \ {v. fst v = 0} \ - {v. snd v = 0} = {0} \ (Sb - {0})" + by (auto simp: zero_prod_def) + show "\S u. finite S \ S \ ?prod \ {0} \ (\v\S. u v *\<^sub>R v) = (a, b)" + apply (rule exI[of _ ?S]) + apply (rule exI[of _ "\(v, w). (if w = 0 then ua v else 0) + (if v = 0 then ub w else 0)"]) + using span + apply (simp add: prod_case_unfold setsum_addf if_distrib cond_application_beta setsum_cases prod *) + by (auto simp add: setsum_prod intro!: setsum_mono_zero_cong_left) + qed simp + + show "\d>0. ?b ` {d..} = {0} \ independent (?b ` {.. inj_on ?b {..j\{..R ?b j) = ?b i" (is "?SUM = _") + let ?left = "{.. = (\j\?left - {i}. u (?b j) *\<^sub>R ?b j) + + (\j\?right. u (?b j) *\<^sub>R ?b j)" + using `i < DIM('a)` by (subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong) + also have "\ = (\j\?left - {i}. u (?b_a j, 0) *\<^sub>R (?b_a j, 0)) + + (\j\?right. u (0, ?b_b (j-DIM('a))) *\<^sub>R (0, ?b_b (j-DIM('a))))" + unfolding basis_prod_def by auto + finally have "basis i = (\j\?left - {i}. u (?b_a j, 0) *\<^sub>R ?b_a j)" + by (simp add: setsum_prod) + moreover + note independent_basis[where 'a='a, unfolded independent_eq_inj_on[OF basis_inj]] + note this[rule_format, of i "\v. u (v, 0)"] + ultimately show False using `i < DIM('a)` by auto + next + let ?i = "i - DIM('a)" + assume not: "\ i < DIM('a)" hence "DIM('a) \ i" by auto + hence "?i < DIM('b)" using `i < DIM('b) + DIM('a)` by auto + + have inj_on: "inj_on (\j. j - DIM('a)) {DIM('a)..j. j-DIM('a))`(?right - {i})" + by (auto simp: inj_on_image_set_diff image_minus_const_atLeastLessThan_nat) + + have "(0, basis ?i) = ?SUM" unfolding `?SUM = ?b i` + unfolding basis_prod_def using not `?i < DIM('b)` by auto + also have "\ = (\j\?left. u (?b j) *\<^sub>R ?b j) + + (\j\?right - {i}. u (?b j) *\<^sub>R ?b j)" + using not by (subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong) + also have "\ = (\j\?left. u (?b_a j, 0) *\<^sub>R (?b_a j, 0)) + + (\j\?right - {i}. u (0, ?b_b (j-DIM('a))) *\<^sub>R (0, ?b_b (j-DIM('a))))" + unfolding basis_prod_def by auto + finally have "basis ?i = (\j\{..R ?b_b j)" + unfolding * + by (subst setsum_reindex[OF inj_on[THEN subset_inj_on]]) + (auto simp: setsum_prod) + moreover + note independent_basis[where 'a='b, unfolded independent_eq_inj_on[OF basis_inj]] + note this[rule_format, of ?i "\v. u (0, v)"] + ultimately show False using `?i < DIM('b)` by auto + qed + qed + qed +qed end + +lemma DIM_prod[simp]: "DIM('a \ 'b) = DIM('b::real_basis) + DIM('a::real_basis)" + by (rule dimension_eq) (auto simp: basis_prod_def zero_prod_def basis_eq_0_iff) + +instance prod :: (euclidean_space, euclidean_space) euclidean_space +proof (default, safe) + let ?b = "basis :: nat \ 'a \ 'b" + fix i j assume "i < DIM('a \ 'b)" "j < DIM('a \ 'b)" + thus "?b i \ ?b j = (if i = j then 1 else 0)" + unfolding basis_prod_def by (auto simp: dot_basis) +qed + +instantiation prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space +begin + +definition "x \ (y::('a\'b)) \ (\i'b). x $$ i \ y $$ i)" +definition "x < (y::('a\'b)) \ (\i'b). x $$ i < y $$ i)" + +instance proof qed (auto simp: less_prod_def less_eq_prod_def) +end + + +end