# HG changeset patch # User immler # Date 1387210102 -3600 # Node ID 13f08c876899cf699a00f87a4d4267d80d7c9bdb # Parent 1a2da44c8e7d3f1b2d8b8e060fc8a7989dab8df9 introduced ordered real vector spaces diff -r 1a2da44c8e7d -r 13f08c876899 src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Mon Dec 16 17:08:22 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Mon Dec 16 17:08:22 2013 +0100 @@ -3129,6 +3129,9 @@ subclass ordered_ab_group_add_abs by default (auto simp: eucl_le inner_add_left eucl_abs abs_leI) +subclass ordered_real_vector + by default (auto simp: eucl_le intro!: mult_left_mono mult_right_mono) + subclass lattice by default (auto simp: eucl_inf eucl_sup eucl_le) diff -r 1a2da44c8e7d -r 13f08c876899 src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Mon Dec 16 17:08:22 2013 +0100 +++ b/src/HOL/Real_Vector_Spaces.thy Mon Dec 16 17:08:22 2013 +0100 @@ -431,6 +431,121 @@ "q \ \ \ (\r. P (of_real r)) \ P q" by (rule Reals_cases) auto +subsection {* Ordered real vector spaces *} + +class ordered_real_vector = real_vector + ordered_ab_group_add + + assumes scaleR_left_mono: "x \ y \ 0 \ a \ a *\<^sub>R x \ a *\<^sub>R y" + assumes scaleR_right_mono: "a \ b \ 0 \ x \ a *\<^sub>R x \ b *\<^sub>R x" +begin + +lemma scaleR_mono: + "a \ b \ x \ y \ 0 \ b \ 0 \ x \ a *\<^sub>R x \ b *\<^sub>R y" +apply (erule scaleR_right_mono [THEN order_trans], assumption) +apply (erule scaleR_left_mono, assumption) +done + +lemma scaleR_mono': + "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a *\<^sub>R c \ b *\<^sub>R d" + by (rule scaleR_mono) (auto intro: order.trans) + +end + +lemma scaleR_nonneg_nonneg: "0 \ a \ 0 \ (x::'a::ordered_real_vector) \ 0 \ a *\<^sub>R x" + using scaleR_left_mono [of 0 x a] + by simp + +lemma scaleR_nonneg_nonpos: "0 \ a \ (x::'a::ordered_real_vector) \ 0 \ a *\<^sub>R x \ 0" + using scaleR_left_mono [of x 0 a] by simp + +lemma scaleR_nonpos_nonneg: "a \ 0 \ 0 \ (x::'a::ordered_real_vector) \ a *\<^sub>R x \ 0" + using scaleR_right_mono [of a 0 x] by simp + +lemma split_scaleR_neg_le: "(0 \ a & x \ 0) | (a \ 0 & 0 \ x) \ + a *\<^sub>R (x::'a::ordered_real_vector) \ 0" + by (auto simp add: scaleR_nonneg_nonpos scaleR_nonpos_nonneg) + +lemma le_add_iff1: + fixes c d e::"'a::ordered_real_vector" + shows "a *\<^sub>R e + c \ b *\<^sub>R e + d \ (a - b) *\<^sub>R e + c \ d" + by (simp add: algebra_simps) + +lemma le_add_iff2: + fixes c d e::"'a::ordered_real_vector" + shows "a *\<^sub>R e + c \ b *\<^sub>R e + d \ c \ (b - a) *\<^sub>R e + d" + by (simp add: algebra_simps) + +lemma scaleR_left_mono_neg: + fixes a b::"'a::ordered_real_vector" + shows "b \ a \ c \ 0 \ c *\<^sub>R a \ c *\<^sub>R b" + apply (drule scaleR_left_mono [of _ _ "- c"]) + apply simp_all + done + +lemma scaleR_right_mono_neg: + fixes c::"'a::ordered_real_vector" + shows "b \ a \ c \ 0 \ a *\<^sub>R c \ b *\<^sub>R c" + apply (drule scaleR_right_mono [of _ _ "- c"]) + apply simp_all + done + +lemma scaleR_nonpos_nonpos: "a \ 0 \ (b::'a::ordered_real_vector) \ 0 \ 0 \ a *\<^sub>R b" +using scaleR_right_mono_neg [of a 0 b] by simp + +lemma split_scaleR_pos_le: + fixes b::"'a::ordered_real_vector" + shows "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a *\<^sub>R b" + by (auto simp add: scaleR_nonneg_nonneg scaleR_nonpos_nonpos) + +lemma zero_le_scaleR_iff: + fixes b::"'a::ordered_real_vector" + shows "0 \ a *\<^sub>R b \ 0 < a \ 0 \ b \ a < 0 \ b \ 0 \ a = 0" (is "?lhs = ?rhs") +proof cases + assume "a \ 0" + show ?thesis + proof + assume lhs: ?lhs + { + assume "0 < a" + with lhs have "inverse a *\<^sub>R 0 \ inverse a *\<^sub>R (a *\<^sub>R b)" + by (intro scaleR_mono) auto + hence ?rhs using `0 < a` + by simp + } moreover { + assume "0 > a" + with lhs have "- inverse a *\<^sub>R 0 \ - inverse a *\<^sub>R (a *\<^sub>R b)" + by (intro scaleR_mono) auto + hence ?rhs using `0 > a` + by simp + } ultimately show ?rhs using `a \ 0` by arith + qed (auto simp: not_le `a \ 0` intro!: split_scaleR_pos_le) +qed simp + +lemma scaleR_le_0_iff: + fixes b::"'a::ordered_real_vector" + shows "a *\<^sub>R b \ 0 \ 0 < a \ b \ 0 \ a < 0 \ 0 \ b \ a = 0" + by (insert zero_le_scaleR_iff [of "-a" b]) force + +lemma scaleR_le_cancel_left: + fixes b::"'a::ordered_real_vector" + shows "c *\<^sub>R a \ c *\<^sub>R b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" + by (auto simp add: neq_iff scaleR_left_mono scaleR_left_mono_neg + dest: scaleR_left_mono[where a="inverse c"] scaleR_left_mono_neg[where c="inverse c"]) + +lemma scaleR_le_cancel_left_pos: + fixes b::"'a::ordered_real_vector" + shows "0 < c \ c *\<^sub>R a \ c *\<^sub>R b \ a \ b" + by (auto simp: scaleR_le_cancel_left) + +lemma scaleR_le_cancel_left_neg: + fixes b::"'a::ordered_real_vector" + shows "c < 0 \ c *\<^sub>R a \ c *\<^sub>R b \ b \ a" + by (auto simp: scaleR_le_cancel_left) + +lemma scaleR_left_le_one_le: + fixes x::"'a::ordered_real_vector" and a::real + shows "0 \ x \ a \ 1 \ a *\<^sub>R x \ x" + using scaleR_right_mono[of a 1 x] by simp + subsection {* Real normed vector spaces *}