# HG changeset patch # User hoelzl # Date 1277967669 -7200 # Node ID 2946b8f057df5935eaf5cd00700a29ef251f12b8 # Parent f2c98b8c0c5c9428065b23295d83f7929319aeb0 Instantiate product type as euclidean space. diff -r f2c98b8c0c5c -r 2946b8f057df src/HOL/Library/Product_plus.thy --- a/src/HOL/Library/Product_plus.thy Thu Jul 01 08:13:20 2010 +0200 +++ b/src/HOL/Library/Product_plus.thy Thu Jul 01 09:01:09 2010 +0200 @@ -118,4 +118,7 @@ lemma snd_setsum: "snd (\x\A. f x) = (\x\A. snd (f x))" by (cases "finite A", induct set: finite, simp_all) +lemma setsum_prod: "(\x\A. (f x, g x)) = (\x\A. f x, \x\A. g x)" +by (cases "finite A", induct set: finite) (simp_all add: zero_prod_def) + end diff -r f2c98b8c0c5c -r 2946b8f057df src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Thu Jul 01 08:13:20 2010 +0200 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Thu Jul 01 09:01:09 2010 +0200 @@ -5,9 +5,148 @@ imports Finite_Cartesian_Product Integration begin -(* TODO: real_vector^'n is instance of real_vector *) +instantiation * :: (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 * :: (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 * :: (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 lemma delta_mult_idempotent: "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto) @@ -40,6 +179,9 @@ subsection{* Basic componentwise operations on vectors. *} +lemma dimindex_ge_1:"CARD(_::finite) \ 1" + using one_le_card_finite by auto + instantiation cart :: (times,finite) times begin definition vector_mult_def : "op * \ (\ x y. (\ i. (x$i) * (y$i)))" @@ -426,17 +568,6 @@ let ?b = "basis :: nat \ 'a^'b" let ?b' = "basis :: nat \ 'a" - { fix D :: nat and f :: "nat \ 'c::real_vector" - assume "inj_on f {..i. i < D \ f ` {..i. inj_on f ({..i. finite (f ` {.. (\a u. a < D \ (\i\{..R f i) \ f a)" - unfolding dependent_def span_finite[OF *] - by (auto simp: eq setsum_reindex[OF inj]) } - note independentI = this - have setsum_basis: "\f. (\x\range basis. f (x::'a)) = f 0 + (\i real" assume "j < DIM('a)" let ?x = "j + \' i * DIM('a)" show "(\k\{..R ?b k) \ ?b ?x" @@ -484,7 +615,7 @@ by (subst setsum_reindex) (auto simp: basis_eq_pi' intro!: inj_onI) also have "\ \ ?b ?x $ i" proof - - note independentI[THEN iffD1, OF basis_inj independent_basis, rule_format] + note independent_eq_inj_on[THEN iffD1, OF basis_inj independent_basis, rule_format] note this[of j "\v. u (\ ka::'b. if ka = i then v else (0\'a))"] thus ?thesis by (simp add: `j < DIM('a)` basis_eq pi'_range) qed @@ -506,7 +637,7 @@ by (rule the_equality) (simp_all split: split_if_asm add: basis_neq_0) { fix x :: 'a have "x \ span (range basis)" - using span_basis by (auto simp: basis_range) + using span_basis by (auto simp: range_basis) hence "\u. (\xR ?b' x) = x" by (subst (asm) span_finite) (auto simp: setsum_basis) } hence "\i. \u. (\xR ?b' x) = i" by auto @@ -523,7 +654,7 @@ by (auto simp: image_def basis_cart_def) ultimately show ?thesis - by (auto simp add: Cart_eq setsum_reindex[OF inj] basis_range + by (auto simp add: Cart_eq setsum_reindex[OF inj] range_basis setsum_mult_product basis_eq if_distrib setsum_cases span_finite setsum_reindex[OF basis_inj]) qed diff -r f2c98b8c0c5c -r 2946b8f057df src/HOL/Multivariate_Analysis/Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Thu Jul 01 08:13:20 2010 +0200 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Thu Jul 01 09:01:09 2010 +0200 @@ -1402,6 +1402,42 @@ ultimately show ?thesis by blast qed +lemma Int_Un_cancel: "(A \ B) \ A = A" "(A \ B) \ B = B" by auto + +lemma span_union: "span (A \ B) = (\(a, b). a + b) ` (span A \ span B)" +proof safe + fix x assume "x \ span (A \ B)" + then obtain S u where S: "finite S" "S \ A \ B" and x: "x = (\v\S. u v *\<^sub>R v)" + unfolding span_explicit by auto + + let ?Sa = "\v\S\A. u v *\<^sub>R v" + let ?Sb = "(\v\S\(B - A). u v *\<^sub>R v)" + show "x \ (\(a, b). a + b) ` (span A \ span B)" + proof + show "x = (case (?Sa, ?Sb) of (a, b) \ a + b)" + unfolding x using S + by (simp, subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong) + + from S have "?Sa \ span A" unfolding span_explicit + by (auto intro!: exI[of _ "S \ A"]) + moreover from S have "?Sb \ span B" unfolding span_explicit + by (auto intro!: exI[of _ "S \ (B - A)"]) + ultimately show "(?Sa, ?Sb) \ span A \ span B" by simp + qed +next + fix a b assume "a \ span A" and "b \ span B" + then obtain Sa ua Sb ub where span: + "finite Sa" "Sa \ A" "a = (\v\Sa. ua v *\<^sub>R v)" + "finite Sb" "Sb \ B" "b = (\v\Sb. ub v *\<^sub>R v)" + unfolding span_explicit by auto + let "?u v" = "(if v \ Sa then ua v else 0) + (if v \ Sb then ub v else 0)" + from span have "finite (Sa \ Sb)" "Sa \ Sb \ A \ B" + and "a + b = (\v\(Sa\Sb). ?u v *\<^sub>R v)" + unfolding setsum_addf scaleR_left_distrib + by (auto simp add: if_distrib cond_application_beta setsum_cases Int_Un_cancel) + thus "a + b \ span (A \ B)" + unfolding span_explicit by (auto intro!: exI[of _ ?u]) +qed text {* This is useful for building a basis step-by-step. *} @@ -1645,10 +1681,6 @@ thus "basis j = 0 \ DIM('a) \ j" unfolding not_le[symmetric] by blast qed -lemma (in real_basis) basis_range: - "range (basis) = {0} \ basis ` {..'a) ` {..'a) ` {.. span (basis ` {.. 'c::real_vector" assumes *: "inj_on f {.. (\a u. a < D \ (\i\{..R f i) \ f a)" +proof - + from * have eq: "\i. i < D \ f ` {..i. inj_on f ({..i. finite (f ` {..'a) ` {..'a) ` {.. finite S \ card S <= DIM('a::euclidean_space)" diff -r f2c98b8c0c5c -r 2946b8f057df src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Thu Jul 01 08:13:20 2010 +0200 +++ b/src/HOL/SetInterval.thy Thu Jul 01 09:01:09 2010 +0200 @@ -456,6 +456,23 @@ apply auto done +lemma image_minus_const_atLeastLessThan_nat: + fixes c :: nat + shows "(\i. i - c) ` {x ..< y} = + (if c < y then {x - c ..< y - c} else if x < y then {0} else {})" + (is "_ = ?right") +proof safe + fix a assume a: "a \ ?right" + show "a \ (\i. i - c) ` {x ..< y}" + proof cases + assume "c < y" with a show ?thesis + by (auto intro!: image_eqI[of _ _ "a + c"]) + next + assume "\ c < y" with a show ?thesis + by (auto intro!: image_eqI[of _ _ x] split: split_if_asm) + qed +qed auto + context ordered_ab_group_add begin