diff -r 2a882ef2cd73 -r 99cf6e470816 src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Mon Dec 19 14:41:08 2011 +0100 +++ b/src/HOL/Orderings.thy Mon Dec 19 14:41:08 2011 +0100 @@ -1050,33 +1050,20 @@ end -lemma min_leastL: "(!!x. least <= x) ==> min least x = least" +lemma min_absorb1: "x \ y \ min x y = x" by (simp add: min_def) -lemma max_leastL: "(!!x. least <= x) ==> max least x = x" +lemma max_absorb2: "x \ y ==> max x y = y" by (simp add: max_def) -lemma min_leastR: "(\x\'a\order. least \ x) \ min x least = least" -by (simp add: min_def) (blast intro: antisym) - -lemma max_leastR: "(\x\'a\order. least \ x) \ max x least = x" -by (simp add: max_def) (blast intro: antisym) - -lemma min_greatestL: "(\x::'a::order. x \ greatest) \ min greatest x = x" -by (simp add: min_def) (blast intro: antisym) +lemma min_absorb2: "(y\'a\order) \ x \ min x y = y" +by (simp add:min_def) -lemma max_greatestL: "(\x::'a::order. x \ greatest) \ max greatest x = greatest" -by (simp add: max_def) (blast intro: antisym) - -lemma min_greatestR: "(\x. x \ greatest) \ min x greatest = x" -by (simp add: min_def) - -lemma max_greatestR: "(\x. x \ greatest) \ max x greatest = greatest" +lemma max_absorb1: "(y\'a\order) \ x \ max x y = x" by (simp add: max_def) - subsection {* (Unique) top and bottom elements *} class bot = order +