# HG changeset patch # User hoelzl # Date 1396260995 -7200 # Node ID b3551501424ea0ced6897336d51edd883065c08a # Parent 3e62e68fb342cce5d42dc266d3d3741322e375ea add rules about infinity of intervals diff -r 3e62e68fb342 -r b3551501424e src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Sun Mar 30 21:24:59 2014 +0200 +++ b/src/HOL/Set_Interval.thy Mon Mar 31 12:16:35 2014 +0200 @@ -414,6 +414,11 @@ using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) +lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff: + "{a <..< b} \ { c <.. d } \ (a < b \ c \ a \ b \ d)" + using dense[of "a" "min c b"] dense[of "max a d" "b"] + by (force simp: subset_eq Ball_def not_less[symmetric]) + end context no_top @@ -463,6 +468,88 @@ by (auto simp: set_eq_iff intro: top_le le_bot) +subsection {* Infinite intervals *} + +context dense_linorder +begin + +lemma infinite_Ioo: + assumes "a < b" + shows "\ finite {a<.. {}" + using `a < b` by auto + ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b" + using Max_in[of "{a <..< b}"] by auto + then obtain x where "Max {a <..< b} < x" "x < b" + using dense[of "Max {a<.. {a <..< b}" + using `a < Max {a <..< b}` by auto + then have "x \ Max {a <..< b}" + using fin by auto + with `Max {a <..< b} < x` show False by auto +qed + +lemma infinite_Icc: "a < b \ \ finite {a .. b}" + using greaterThanLessThan_subseteq_atLeastAtMost_iff[of a b a b] infinite_Ioo[of a b] + by (auto dest: finite_subset) + +lemma infinite_Ico: "a < b \ \ finite {a ..< b}" + using greaterThanLessThan_subseteq_atLeastLessThan_iff[of a b a b] infinite_Ioo[of a b] + by (auto dest: finite_subset) + +lemma infinite_Ioc: "a < b \ \ finite {a <.. b}" + using greaterThanLessThan_subseteq_greaterThanAtMost_iff[of a b a b] infinite_Ioo[of a b] + by (auto dest: finite_subset) + +end + +lemma infinite_Iio: "\ finite {..< a :: 'a :: {no_bot, linorder}}" +proof + assume "finite {..< a}" + then have *: "\x. x < a \ Min {..< a} \ x" + by auto + obtain x where "x < a" + using lt_ex by auto + + obtain y where "y < Min {..< a}" + using lt_ex by auto + also have "Min {..< a} \ x" + using `x < a` by fact + also note `x < a` + finally have "Min {..< a} \ y" + by fact + with `y < Min {..< a}` show False by auto +qed + +lemma infinite_Iic: "\ finite {.. a :: 'a :: {no_bot, linorder}}" + using infinite_Iio[of a] finite_subset[of "{..< a}" "{.. a}"] + by (auto simp: subset_eq less_imp_le) + +lemma infinite_Ioi: "\ finite {a :: 'a :: {no_top, linorder} <..}" +proof + assume "finite {a <..}" + then have *: "\x. a < x \ x \ Max {a <..}" + by auto + + obtain y where "Max {a <..} < y" + using gt_ex by auto + + obtain x where "a < x" + using gt_ex by auto + also then have "x \ Max {a <..}" + by fact + also note `Max {a <..} < y` + finally have "y \ Max { a <..}" + by fact + with `Max {a <..} < y` show False by auto +qed + +lemma infinite_Ici: "\ finite {a :: 'a :: {no_top, linorder} ..}" + using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"] + by (auto simp: subset_eq less_imp_le) + subsubsection {* Intersection *} context linorder @@ -823,6 +910,7 @@ "!!f::nat=>nat. (!!n. n \ f n) ==> finite {n. f n \ u}" by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) + text{* Any subset of an interval of natural numbers the size of the subset is exactly that interval. *}