diff -r 0d399131008f -r d50b993b4fb9 src/HOL/Library/Infinite_Set.thy --- a/src/HOL/Library/Infinite_Set.thy Mon Nov 30 14:24:51 2015 +0100 +++ b/src/HOL/Library/Infinite_Set.thy Tue Dec 01 14:09:10 2015 +0000 @@ -72,10 +72,7 @@ by(induction rule: finite_psubset_induct)(meson Diff_subset card_Diff1_less card_psubset finite_Diff step) qed -text \ - As a concrete example, we prove that the set of natural numbers is - infinite. -\ +text \As a concrete example, we prove that the set of natural numbers is infinite.\ lemma infinite_nat_iff_unbounded_le: "infinite (S::nat set) \ (\m. \n\m. n \ S)" using frequently_cofinite[of "\x. x \ S"] @@ -94,6 +91,7 @@ lemma finite_nat_bounded: "finite (S::nat set) \ \k. S \ {.. For a set of natural numbers to be infinite, it is enough to know that for any number larger than some \k\, there is some larger @@ -150,6 +148,32 @@ obtains y where "y \ f`A" and "infinite (f -` {y})" using assms by (blast dest: inf_img_fin_dom) +proposition finite_image_absD: + fixes S :: "'a::linordered_ring set" + shows "finite (abs ` S) \ finite S" + by (rule ccontr) (auto simp: abs_eq_iff vimage_def dest: inf_img_fin_dom) + +text \The set of integers is also infinite.\ + +lemma infinite_int_iff_infinite_nat_abs: "infinite (S::int set) \ infinite ((nat o abs) ` S)" + by (auto simp: transfer_nat_int_set_relations o_def image_comp dest: finite_image_absD) + +proposition infinite_int_iff_unbounded_le: "infinite (S::int set) \ (\m. \n. abs n \ m \ n \ S)" + apply (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded_le o_def image_def) + apply (metis abs_ge_zero nat_le_eq_zle le_nat_iff) + done + +proposition infinite_int_iff_unbounded: "infinite (S::int set) \ (\m. \n. abs n > m \ n \ S)" + apply (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded o_def image_def) + apply (metis (full_types) nat_le_iff nat_mono not_le) + done + +proposition finite_int_iff_bounded: "finite (S::int set) \ (\k. abs ` S \ {.. (\k. abs ` S \ {.. k})" + using infinite_int_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) + subsection "Infinitely Many and Almost All" text \ @@ -385,24 +409,5 @@ unfolding bij_betw_def by (auto intro: enumerate_in_set) qed -subsection "Miscellaneous" - -text \ - A few trivial lemmas about sets that contain at most one element. - These simplify the reasoning about deterministic automata. -\ - -definition atmost_one :: "'a set \ bool" - where "atmost_one S \ (\x y. x\S \ y\S \ x = y)" - -lemma atmost_one_empty: "S = {} \ atmost_one S" - by (simp add: atmost_one_def) - -lemma atmost_one_singleton: "S = {x} \ atmost_one S" - by (simp add: atmost_one_def) - -lemma atmost_one_unique [elim]: "atmost_one S \ x \ S \ y \ S \ y = x" - by (simp add: atmost_one_def) - end