# HG changeset patch # User haftmann # Date 1311266431 -7200 # Node ID 481566bc20e43078519660de5212f8843dd8482b # Parent 26ca0bad226ae02060439e285242bb941e16b9ed ereal is a complete_linorder instance diff -r 26ca0bad226a -r 481566bc20e4 src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Wed Jul 20 22:14:39 2011 +0200 +++ b/src/HOL/Library/Extended_Real.thy Thu Jul 21 18:40:31 2011 +0200 @@ -8,7 +8,7 @@ header {* Extended real number line *} theory Extended_Real - imports Complex_Main Extended_Nat +imports Complex_Main Extended_Nat begin text {* @@ -1244,8 +1244,11 @@ with Sup_le[of "uminus`A" "-x"] show "x \ Inf A" unfolding ereal_Sup_uminus_image_eq by force } qed + end +instance ereal :: complete_linorder .. + lemma ereal_SUPR_uminus: fixes f :: "'a => ereal" shows "(SUP i : R. -(f i)) = -(INF i : R. f i)" @@ -1335,12 +1338,11 @@ by (cases e) auto qed -lemma Sup_eq_top_iff: - fixes A :: "'a::{complete_lattice, linorder} set" - shows "Sup A = top \ (\xi\A. x < i)" +lemma (in complete_linorder) Sup_eq_top_iff: -- "FIXME move" + "Sup A = top \ (\xi\A. x < i)" proof assume *: "Sup A = top" - show "(\xi\A. x < i)" unfolding *[symmetric] + show "(\xi\A. x < i)" unfolding * [symmetric] proof (intro allI impI) fix x assume "x < Sup A" then show "\i\A. x < i" unfolding less_Sup_iff by auto @@ -1350,7 +1352,7 @@ show "Sup A = top" proof (rule ccontr) assume "Sup A \ top" - with top_greatest[of "Sup A"] + with top_greatest [of "Sup A"] have "Sup A < top" unfolding le_less by auto then have "Sup A < Sup A" using * unfolding less_Sup_iff by auto @@ -1358,8 +1360,8 @@ qed qed -lemma SUP_eq_top_iff: - fixes f :: "'a \ 'b::{complete_lattice, linorder}" +lemma (in complete_linorder) SUP_eq_top_iff: -- "FIXME move" + fixes f :: "'b \ 'a" shows "(SUP i:A. f i) = top \ (\xi\A. x < f i)" unfolding SUPR_def Sup_eq_top_iff by auto @@ -2182,12 +2184,12 @@ "Limsup net f = (LEAST l. \y>l. eventually (\x. f x < y) net)" lemma Liminf_Sup: - fixes f :: "'a => 'b::{complete_lattice, linorder}" + fixes f :: "'a => 'b::complete_linorder" shows "Liminf net f = Sup {l. \yx. y < f x) net}" by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def) lemma Limsup_Inf: - fixes f :: "'a => 'b::{complete_lattice, linorder}" + fixes f :: "'a => 'b::complete_linorder" shows "Limsup net f = Inf {l. \y>l. eventually (\x. f x < y) net}" by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def) @@ -2208,7 +2210,7 @@ using assms by (auto intro!: Greatest_equality) lemma Limsup_const: - fixes c :: "'a::{complete_lattice, linorder}" + fixes c :: "'a::complete_linorder" assumes ntriv: "\ trivial_limit net" shows "Limsup net (\x. c) = c" unfolding Limsup_Inf @@ -2222,7 +2224,7 @@ qed auto lemma Liminf_const: - fixes c :: "'a::{complete_lattice, linorder}" + fixes c :: "'a::complete_linorder" assumes ntriv: "\ trivial_limit net" shows "Liminf net (\x. c) = c" unfolding Liminf_Sup @@ -2235,18 +2237,17 @@ qed qed auto -lemma mono_set: - fixes S :: "('a::order) set" - shows "mono S \ (\x y. x \ y \ x \ S \ y \ S)" +lemma (in order) mono_set: + "mono S \ (\x y. x \ y \ x \ S \ y \ S)" by (auto simp: mono_def mem_def) -lemma mono_greaterThan[intro, simp]: "mono {B<..}" unfolding mono_set by auto -lemma mono_atLeast[intro, simp]: "mono {B..}" unfolding mono_set by auto -lemma mono_UNIV[intro, simp]: "mono UNIV" unfolding mono_set by auto -lemma mono_empty[intro, simp]: "mono {}" unfolding mono_set by auto +lemma (in order) mono_greaterThan [intro, simp]: "mono {B<..}" unfolding mono_set by auto +lemma (in order) mono_atLeast [intro, simp]: "mono {B..}" unfolding mono_set by auto +lemma (in order) mono_UNIV [intro, simp]: "mono UNIV" unfolding mono_set by auto +lemma (in order) mono_empty [intro, simp]: "mono {}" unfolding mono_set by auto -lemma mono_set_iff: - fixes S :: "'a::{linorder,complete_lattice} set" +lemma (in complete_linorder) mono_set_iff: + fixes S :: "'a set" defines "a \ Inf S" shows "mono S \ (S = {a <..} \ S = {a..})" (is "_ = ?c") proof @@ -2257,13 +2258,13 @@ assume "a \ S" show ?c using mono[OF _ `a \ S`] - by (auto intro: complete_lattice_class.Inf_lower simp: a_def) + by (auto intro: Inf_lower simp: a_def) next assume "a \ S" have "S = {a <..}" proof safe fix x assume "x \ S" - then have "a \ x" unfolding a_def by (rule complete_lattice_class.Inf_lower) + then have "a \ x" unfolding a_def by (rule Inf_lower) then show "a < x" using `x \ S` `a \ S` by (cases "a = x") auto next fix x assume "a < x" diff -r 26ca0bad226a -r 481566bc20e4 src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Wed Jul 20 22:14:39 2011 +0200 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Thu Jul 21 18:40:31 2011 +0200 @@ -6,7 +6,7 @@ header {*Lebesgue Integration*} theory Lebesgue_Integration -imports Measure Borel_Space + imports Measure Borel_Space begin lemma real_ereal_1[simp]: "real (1::ereal) = 1"