# HG changeset patch # User Andreas Lochbihler # Date 1429004657 -7200 # Node ID 3630ecde4e7c1df296e7b09d836ee71348c8a763 # Parent 8a6d947cc756cdd6c790c0c0e463dd29cd536c6d more lemmas about ereal diff -r 8a6d947cc756 -r 3630ecde4e7c src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Tue Apr 14 11:36:03 2015 +0200 +++ b/src/HOL/Library/Extended_Real.thy Tue Apr 14 11:44:17 2015 +0200 @@ -298,6 +298,10 @@ end +lemma ereal_0_plus [simp]: "ereal 0 + x = x" + and plus_ereal_0 [simp]: "x + ereal 0 = x" +by(simp_all add: zero_ereal_def[symmetric]) + instance ereal :: numeral .. lemma real_of_ereal_0[simp]: "real (0::ereal) = 0" @@ -1286,6 +1290,9 @@ shows "real (a - b) = (if \a\ = \ \ \b\ = \ then 0 else real a - real b)" by (cases rule: ereal2_cases[of a b]) auto +lemma real_of_ereal_minus': "\x\ = \ \ \y\ = \ \ real x - real y = real (x - y :: ereal)" +by(subst real_of_ereal_minus) auto + lemma ereal_diff_positive: fixes a b :: ereal shows "a \ b \ 0 \ b - a" by (cases rule: ereal2_cases[of a b]) auto @@ -1428,6 +1435,14 @@ shows "0 < inverse x \ x \ \ \ 0 \ x" by (cases x) auto +lemma ereal_inverse_le_0_iff: + fixes x :: ereal + shows "inverse x \ 0 \ x < 0 \ x = \" + by(cases x) auto + +lemma ereal_divide_eq_0_iff: "x / y = 0 \ x = 0 \ \y :: ereal\ = \" +by(cases x y rule: ereal2_cases) simp_all + lemma ereal_mult_less_right: fixes a b c :: ereal assumes "b * a < c * a" @@ -1971,7 +1986,7 @@ assumes f: "\i. i \ I \ 0 \ f i" and c: "0 \ c" shows "(SUP i:I. c * f i) = c * (SUP i:I. f i)" proof cases - assume "(SUP i:I. f i) = 0" + assume "(SUP i: I. f i) = 0" moreover then have "\i. i \ I \ f i = 0" by (metis SUP_upper f antisym) ultimately show ?thesis @@ -2646,6 +2661,38 @@ finally show "(SUP P:?F. ?INF P f + (SUP P:?F. ?INF P g)) \ (SUP P:?F. INF x:Collect P. f x + g x)" . qed +lemma Sup_ereal_mult_right': + assumes nonempty: "Y \ {}" + and x: "x \ 0" + shows "(SUP i:Y. f i) * ereal x = (SUP i:Y. f i * ereal x)" (is "?lhs = ?rhs") +proof(cases "x = 0") + case True thus ?thesis by(auto simp add: nonempty zero_ereal_def[symmetric]) +next + case False + show ?thesis + proof(rule antisym) + show "?rhs \ ?lhs" + by(rule SUP_least)(simp add: ereal_mult_right_mono SUP_upper x) + next + have "?lhs / ereal x = (SUP i:Y. f i) * (ereal x / ereal x)" by(simp only: ereal_times_divide_eq) + also have "\ = (SUP i:Y. f i)" using False by simp + also have "\ \ ?rhs / x" + proof(rule SUP_least) + fix i + assume "i \ Y" + have "f i = f i * (ereal x / ereal x)" using False by simp + also have "\ = f i * x / x" by(simp only: ereal_times_divide_eq) + also from \i \ Y\ have "f i * x \ ?rhs" by(rule SUP_upper) + hence "f i * x / x \ ?rhs / x" using x False by simp + finally show "f i \ ?rhs / x" . + qed + finally have "(?lhs / x) * x \ (?rhs / x) * x" + by(rule ereal_mult_right_mono)(simp add: x) + also have "\ = ?rhs" using False ereal_divide_eq mult.commute by force + also have "(?lhs / x) * x = ?lhs" using False ereal_divide_eq mult.commute by force + finally show "?lhs \ ?rhs" . + qed +qed subsubsection {* Tests for code generator *}