diff -r d53d7ca3303e -r 1ab49aa7f3c0 src/HOL/Library/Infinite_Set.thy --- a/src/HOL/Library/Infinite_Set.thy Sun Jan 29 17:27:02 2017 +0100 +++ b/src/HOL/Library/Infinite_Set.thy Mon Jan 30 16:10:52 2017 +0100 @@ -5,26 +5,31 @@ section \Infinite Sets and Related Concepts\ theory Infinite_Set -imports Main + imports Main begin -text \The set of natural numbers is infinite.\ +subsection \The set of natural numbers is infinite\ -lemma infinite_nat_iff_unbounded_le: "infinite (S::nat set) \ (\m. \n\m. n \ S)" +lemma infinite_nat_iff_unbounded_le: "infinite S \ (\m. \n\m. n \ S)" + for S :: "nat set" using frequently_cofinite[of "\x. x \ S"] by (simp add: cofinite_eq_sequentially frequently_def eventually_sequentially) -lemma infinite_nat_iff_unbounded: "infinite (S::nat set) \ (\m. \n>m. n \ S)" +lemma infinite_nat_iff_unbounded: "infinite S \ (\m. \n>m. n \ S)" + for S :: "nat set" using frequently_cofinite[of "\x. x \ S"] by (simp add: cofinite_eq_sequentially frequently_def eventually_at_top_dense) -lemma finite_nat_iff_bounded: "finite (S::nat set) \ (\k. S \ {.. (\k. S \ {.. (\k. S \ {.. k})" +lemma finite_nat_iff_bounded_le: "finite S \ (\k. S \ {.. k})" + for S :: "nat set" using infinite_nat_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) -lemma finite_nat_bounded: "finite (S::nat set) \ \k. S \ {.. \k. S \ {.. lemma unbounded_k_infinite: "\m>k. \n>m. n \ S \ infinite (S::nat set)" -apply (clarsimp simp add: finite_nat_set_iff_bounded) -apply (drule_tac x="Suc (max m k)" in spec) -using less_Suc_eq by fastforce + apply (clarsimp simp add: finite_nat_set_iff_bounded) + apply (drule_tac x="Suc (max m k)" in spec) + using less_Suc_eq apply fastforce + done lemma nat_not_finite: "finite (UNIV::nat set) \ R" by simp lemma range_inj_infinite: - "inj (f::nat \ 'a) \ infinite (range f)" + fixes f :: "nat \ 'a" + assumes "inj f" + shows "infinite (range f)" proof - assume "finite (range f)" and "inj f" - then have "finite (UNIV::nat set)" + assume "finite (range f)" + from this assms have "finite (UNIV::nat set)" by (rule finite_imageD) then show False by simp qed -text \The set of integers is also infinite.\ + +subsection \The set of integers is also infinite\ -lemma infinite_int_iff_infinite_nat_abs: "infinite (S::int set) \ infinite ((nat o abs) ` S)" +lemma infinite_int_iff_infinite_nat_abs: "infinite S \ infinite ((nat \ abs) ` S)" + for S :: "int set" 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. \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_le: "infinite S \ (\m. \n. \n\ \ m \ n \ S)" + for S :: "int set" + by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded_le o_def image_def) + (metis abs_ge_zero nat_le_eq_zle le_nat_iff) -proposition infinite_int_iff_unbounded: "infinite (S::int set) \ (\m. \n. \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 infinite_int_iff_unbounded: "infinite S \ (\m. \n. \n\ > m \ n \ S)" + for S :: "int set" + by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded o_def image_def) + (metis (full_types) nat_le_iff nat_mono not_le) -proposition finite_int_iff_bounded: "finite (S::int set) \ (\k. abs ` S \ {.. (\k. abs ` S \ {.. (\k. abs ` S \ {.. k})" +proposition finite_int_iff_bounded_le: "finite S \ (\k. abs ` S \ {.. k})" + for S :: "int set" using infinite_int_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) -subsection "Infinitely Many and Almost All" + +subsection \Infinitely Many and Almost All\ text \ We often need to reason about the existence of infinitely many @@ -80,9 +93,11 @@ we introduce corresponding binders and their proof rules. \ -(* The following two lemmas are available as filter-rules, but not in the simp-set *) -lemma not_INFM [simp]: "\ (INFM x. P x) \ (MOST x. \ P x)" by (fact not_frequently) -lemma not_MOST [simp]: "\ (MOST x. P x) \ (INFM x. \ P x)" by (fact not_eventually) +lemma not_INFM [simp]: "\ (INFM x. P x) \ (MOST x. \ P x)" + by (rule not_frequently) + +lemma not_MOST [simp]: "\ (MOST x. P x) \ (INFM x. \ P x)" + by (rule not_eventually) lemma INFM_const [simp]: "(INFM x::'a. P) \ P \ infinite (UNIV::'a set)" by (simp add: frequently_const_iff) @@ -91,7 +106,7 @@ by (simp add: eventually_const_iff) lemma INFM_imp_distrib: "(INFM x. P x \ Q x) \ ((MOST x. P x) \ (INFM x. Q x))" - by (simp only: imp_conv_disj frequently_disj_iff not_eventually) + by (rule frequently_imp_iff) lemma MOST_imp_iff: "MOST x. P x \ (MOST x. P x \ Q x) \ (MOST x. Q x)" by (auto intro: eventually_rev_mp eventually_mono) @@ -99,6 +114,7 @@ lemma INFM_conjI: "INFM x. P x \ MOST x. Q x \ INFM x. P x \ Q x" by (rule frequently_rev_mp[of P]) (auto elim: eventually_mono) + text \Properties of quantifiers with injective functions.\ lemma INFM_inj: "INFM x. P (f x) \ inj f \ INFM x. P x" @@ -107,6 +123,7 @@ lemma MOST_inj: "MOST x. P x \ inj f \ MOST x. P (f x)" using finite_vimageI[of "{x. \ P x}" f] by (auto simp: eventually_cofinite) + text \Properties of quantifiers with singletons.\ lemma not_INFM_eq [simp]: @@ -134,18 +151,23 @@ "MOST x. a = x \ P x" unfolding eventually_cofinite by simp_all + text \Properties of quantifiers over the naturals.\ -lemma MOST_nat: "(\\<^sub>\n. P (n::nat)) \ (\m. \n>m. P n)" +lemma MOST_nat: "(\\<^sub>\n. P n) \ (\m. \n>m. P n)" + for P :: "nat \ bool" by (auto simp add: eventually_cofinite finite_nat_iff_bounded_le subset_eq not_le[symmetric]) -lemma MOST_nat_le: "(\\<^sub>\n. P (n::nat)) \ (\m. \n\m. P n)" +lemma MOST_nat_le: "(\\<^sub>\n. P n) \ (\m. \n\m. P n)" + for P :: "nat \ bool" by (auto simp add: eventually_cofinite finite_nat_iff_bounded subset_eq not_le[symmetric]) -lemma INFM_nat: "(\\<^sub>\n. P (n::nat)) \ (\m. \n>m. P n)" +lemma INFM_nat: "(\\<^sub>\n. P n) \ (\m. \n>m. P n)" + for P :: "nat \ bool" by (simp add: frequently_cofinite infinite_nat_iff_unbounded) -lemma INFM_nat_le: "(\\<^sub>\n. P (n::nat)) \ (\m. \n\m. P n)" +lemma INFM_nat_le: "(\\<^sub>\n. P n) \ (\m. \n\m. P n)" + for P :: "nat \ bool" by (simp add: frequently_cofinite infinite_nat_iff_unbounded_le) lemma MOST_INFM: "infinite (UNIV::'a set) \ MOST x::'a. P x \ INFM x::'a. P x" @@ -154,9 +176,8 @@ lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \ (MOST n. P n)" by (simp add: cofinite_eq_sequentially) -lemma - shows MOST_SucI: "MOST n. P n \ MOST n. P (Suc n)" - and MOST_SucD: "MOST n. P (Suc n) \ MOST n. P n" +lemma MOST_SucI: "MOST n. P n \ MOST n. P (Suc n)" + and MOST_SucD: "MOST n. P (Suc n) \ MOST n. P n" by (simp_all add: MOST_Suc_iff) lemma MOST_ge_nat: "MOST n::nat. m \ n" @@ -181,31 +202,34 @@ lemma MOST_I: "(\x. P x) \ MOST x. P x" by (rule eventuallyI) lemmas MOST_iff_finiteNeg = MOST_iff_cofinite -subsection "Enumeration of an Infinite Set" -text \ - The set's element type must be wellordered (e.g. the natural numbers). -\ +subsection \Enumeration of an Infinite Set\ + +text \The set's element type must be wellordered (e.g. the natural numbers).\ text \ Could be generalized to - @{term "enumerate' S n = (SOME t. t \ s \ finite {s\S. s < t} \ card {s\S. s < t} = n)"}. + @{prop "enumerate' S n = (SOME t. t \ s \ finite {s\S. s < t} \ card {s\S. s < t} = n)"}. \ primrec (in wellorder) enumerate :: "'a set \ nat \ 'a" -where - enumerate_0: "enumerate S 0 = (LEAST n. n \ S)" -| enumerate_Suc: "enumerate S (Suc n) = enumerate (S - {LEAST n. n \ S}) n" + where + enumerate_0: "enumerate S 0 = (LEAST n. n \ S)" + | enumerate_Suc: "enumerate S (Suc n) = enumerate (S - {LEAST n. n \ S}) n" lemma enumerate_Suc': "enumerate S (Suc n) = enumerate (S - {enumerate S 0}) n" by simp lemma enumerate_in_set: "infinite S \ enumerate S n \ S" - apply (induct n arbitrary: S) - apply (fastforce intro: LeastI dest!: infinite_imp_nonempty) - apply simp - apply (metis DiffE infinite_remove) - done +proof (induct n arbitrary: S) + case 0 + then show ?case + by (fastforce intro: LeastI dest!: infinite_imp_nonempty) +next + case (Suc n) + then show ?case + by simp (metis DiffE infinite_remove) +qed declare enumerate_0 [simp del] enumerate_Suc [simp del] @@ -221,10 +245,7 @@ done lemma enumerate_mono: "m < n \ infinite S \ enumerate S m < enumerate S n" - apply (erule less_Suc_induct) - apply (auto intro: enumerate_step) - done - + by (induct m n rule: less_Suc_induct) (auto intro: enumerate_step) lemma le_enumerate: assumes S: "infinite S" @@ -258,23 +279,25 @@ using enumerate_mono[OF zero_less_Suc \infinite S\, of n] \infinite S\ apply (subst (1 2) enumerate_Suc') apply (subst Suc) - using \infinite S\ - apply simp + apply (use \infinite S\ in simp) apply (intro arg_cong[where f = Least] ext) apply (auto simp: enumerate_Suc'[symmetric]) done qed lemma enumerate_Ex: - assumes S: "infinite (S::nat set)" - shows "s \ S \ \n. enumerate S n = s" + fixes S :: "nat set" + assumes S: "infinite S" + and s: "s \ S" + shows "\n. enumerate S n = s" + using s proof (induct s rule: less_induct) case (less s) show ?case - proof cases + proof (cases "\y\S. y < s") + case True let ?y = "Max {s'\S. s' < s}" - assume "\y\S. y < s" - then have y: "\x. ?y < x \ (\s'\S. s' < s \ s' < x)" + from True have y: "\x. ?y < x \ (\s'\S. s' < s \ s' < x)" by (subst Max_less_iff) auto then have y_in: "?y \ {s'\S. s' < s}" by (intro Max_in) auto @@ -282,9 +305,9 @@ by auto with S have "enumerate S (Suc n) = s" by (auto simp: y less enumerate_Suc'' intro!: Least_equality) - then show ?case by auto + then show ?thesis by auto next - assume *: "\ (\y\S. y < s)" + case False then have "\t\S. s \ t" by auto with \s \ S\ show ?thesis by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) @@ -307,36 +330,41 @@ unfolding bij_betw_def by (auto intro: enumerate_in_set) qed -text\A pair of weird and wonderful lemmas from HOL Light\ +text \A pair of weird and wonderful lemmas from HOL Light.\ lemma finite_transitivity_chain: assumes "finite A" - and R: "\x. ~ R x x" "\x y z. \R x y; R y z\ \ R x z" - and A: "\x. x \ A \ \y. y \ A \ R x y" + and R: "\x. \ R x x" "\x y z. \R x y; R y z\ \ R x z" + and A: "\x. x \ A \ \y. y \ A \ R x y" shows "A = {}" -using \finite A\ A -proof (induction A) + using \finite A\ A +proof (induct A) + case empty + then show ?case by simp +next case (insert a A) with R show ?case - by (metis empty_iff insert_iff) -qed simp + by (metis empty_iff insert_iff) (* somewhat slow *) +qed corollary Union_maximal_sets: assumes "finite \" - shows "\{T \ \. \U\\. \ T \ U} = \\" + shows "\{T \ \. \U\\. \ T \ U} = \\" (is "?lhs = ?rhs") proof + show "?lhs \ ?rhs" by force show "?rhs \ ?lhs" proof (rule Union_subsetI) fix S assume "S \ \" - have "{T \ \. S \ T} = {}" if "~ (\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y)" + have "{T \ \. S \ T} = {}" + if "\ (\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y)" apply (rule finite_transitivity_chain [of _ "\T U. S \ T \ T \ U"]) - using assms that apply auto - by (blast intro: dual_order.trans psubset_imp_subset) - then show "\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y" - using \S \ \\ by blast + apply (use assms that in auto) + apply (blast intro: dual_order.trans psubset_imp_subset) + done + with \S \ \\ show "\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y" + by blast qed -qed force +qed end -