Up to index of Isabelle/HOL/HOL-Algebra/BPL
theory Infinite_Set(* Title: HOL/Infinite_Set.thy ID: $Id: Infinite_Set.thy,v 1.7 2007/06/14 21:04:39 wenzelm Exp $ Author: Stephan Merz *) header {* Infinite Sets and Related Concepts *} theory Infinite_Set imports Main begin subsection "Infinite Sets" text {* Some elementary facts about infinite sets, mostly by Stefan Merz. Beware! Because "infinite" merely abbreviates a negation, these lemmas may not work well with @{text "blast"}. *} abbreviation infinite :: "'a set => bool" where "infinite S == ¬ finite S" text {* Infinite sets are non-empty, and if we remove some elements from an infinite set, the result is still infinite. *} lemma infinite_imp_nonempty: "infinite S ==> S ≠ {}" by auto lemma infinite_remove: "infinite S ==> infinite (S - {a})" by simp lemma Diff_infinite_finite: assumes T: "finite T" and S: "infinite S" shows "infinite (S - T)" using T proof induct from S show "infinite (S - {})" by auto next fix T x assume ih: "infinite (S - T)" have "S - (insert x T) = (S - T) - {x}" by (rule Diff_insert) with ih show "infinite (S - (insert x T))" by (simp add: infinite_remove) qed lemma Un_infinite: "infinite S ==> infinite (S ∪ T)" by simp lemma infinite_super: assumes T: "S ⊆ T" and S: "infinite S" shows "infinite T" proof assume "finite T" with T have "finite S" by (simp add: finite_subset) with S show False by simp qed text {* As a concrete example, we prove that the set of natural numbers is infinite. *} lemma finite_nat_bounded: assumes S: "finite (S::nat set)" shows "∃k. S ⊆ {..<k}" (is "∃k. ?bounded S k") using S proof induct have "?bounded {} 0" by simp then show "∃k. ?bounded {} k" .. next fix S x assume "∃k. ?bounded S k" then obtain k where k: "?bounded S k" .. show "∃k. ?bounded (insert x S) k" proof (cases "x < k") case True with k show ?thesis by auto next case False with k have "?bounded S (Suc x)" by auto then show ?thesis by auto qed qed lemma finite_nat_iff_bounded: "finite (S::nat set) = (∃k. S ⊆ {..<k})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (rule finite_nat_bounded) next assume ?rhs then obtain k where "S ⊆ {..<k}" .. then show "finite S" by (rule finite_subset) simp qed lemma finite_nat_iff_bounded_le: "finite (S::nat set) = (∃k. S ⊆ {..k})" (is "?lhs = ?rhs") proof assume ?lhs then obtain k where "S ⊆ {..<k}" by (blast dest: finite_nat_bounded) then have "S ⊆ {..k}" by auto then show ?rhs .. next assume ?rhs then obtain k where "S ⊆ {..k}" .. then show "finite S" by (rule finite_subset) simp qed lemma infinite_nat_iff_unbounded: "infinite (S::nat set) = (∀m. ∃n. m<n ∧ n∈S)" (is "?lhs = ?rhs") proof assume ?lhs show ?rhs proof (rule ccontr) assume "¬ ?rhs" then obtain m where m: "∀n. m<n --> n∉S" by blast then have "S ⊆ {..m}" by (auto simp add: sym [OF linorder_not_less]) with `?lhs` show False by (simp add: finite_nat_iff_bounded_le) qed next assume ?rhs show ?lhs proof assume "finite S" then obtain m where "S ⊆ {..m}" by (auto simp add: finite_nat_iff_bounded_le) then have "∀n. m<n --> n∉S" by auto with `?rhs` show False by blast qed qed lemma infinite_nat_iff_unbounded_le: "infinite (S::nat set) = (∀m. ∃n. m≤n ∧ n∈S)" (is "?lhs = ?rhs") proof assume ?lhs show ?rhs proof fix m from `?lhs` obtain n where "m<n ∧ n∈S" by (auto simp add: infinite_nat_iff_unbounded) then have "m≤n ∧ n∈S" by simp then show "∃n. m ≤ n ∧ n ∈ S" .. qed next assume ?rhs show ?lhs proof (auto simp add: infinite_nat_iff_unbounded) fix m from `?rhs` obtain n where "Suc m ≤ n ∧ n∈S" by blast then have "m<n ∧ n∈S" by simp then show "∃n. m < n ∧ n ∈ S" .. qed qed text {* For a set of natural numbers to be infinite, it is enough to know that for any number larger than some @{text k}, there is some larger number that is an element of the set. *} lemma unbounded_k_infinite: assumes k: "∀m. k<m --> (∃n. m<n ∧ n∈S)" shows "infinite (S::nat set)" proof - { fix m have "∃n. m<n ∧ n∈S" proof (cases "k<m") case True with k show ?thesis by blast next case False from k obtain n where "Suc k < n ∧ n∈S" by auto with False have "m<n ∧ n∈S" by auto then show ?thesis .. qed } then show ?thesis by (auto simp add: infinite_nat_iff_unbounded) qed lemma nat_infinite [simp]: "infinite (UNIV :: nat set)" by (auto simp add: infinite_nat_iff_unbounded) lemma nat_not_finite [elim]: "finite (UNIV::nat set) ==> R" by simp text {* Every infinite set contains a countable subset. More precisely we show that a set @{text S} is infinite if and only if there exists an injective function from the naturals into @{text S}. *} lemma range_inj_infinite: "inj (f::nat => 'a) ==> infinite (range f)" proof assume "inj f" and "finite (range f)" then have "finite (UNIV::nat set)" by (auto intro: finite_imageD simp del: nat_infinite) then show False by simp qed lemma int_infinite [simp]: shows "infinite (UNIV::int set)" proof - from inj_int have "infinite (range int)" by (rule range_inj_infinite) moreover have "range int ⊆ (UNIV::int set)" by simp ultimately show "infinite (UNIV::int set)" by (simp add: infinite_super) qed text {* The ``only if'' direction is harder because it requires the construction of a sequence of pairwise different elements of an infinite set @{text S}. The idea is to construct a sequence of non-empty and infinite subsets of @{text S} obtained by successively removing elements of @{text S}. *} lemma linorder_injI: assumes hyp: "!!x y. x < (y::'a::linorder) ==> f x ≠ f y" shows "inj f" proof (rule inj_onI) fix x y assume f_eq: "f x = f y" show "x = y" proof (rule linorder_cases) assume "x < y" with hyp have "f x ≠ f y" by blast with f_eq show ?thesis by simp next assume "x = y" then show ?thesis . next assume "y < x" with hyp have "f y ≠ f x" by blast with f_eq show ?thesis by simp qed qed lemma infinite_countable_subset: assumes inf: "infinite (S::'a set)" shows "∃f. inj (f::nat => 'a) ∧ range f ⊆ S" proof - def Sseq ≡ "nat_rec S (λn T. T - {SOME e. e ∈ T})" def pick ≡ "λn. (SOME e. e ∈ Sseq n)" have Sseq_inf: "!!n. infinite (Sseq n)" proof - fix n show "infinite (Sseq n)" proof (induct n) from inf show "infinite (Sseq 0)" by (simp add: Sseq_def) next fix n assume "infinite (Sseq n)" then show "infinite (Sseq (Suc n))" by (simp add: Sseq_def infinite_remove) qed qed have Sseq_S: "!!n. Sseq n ⊆ S" proof - fix n show "Sseq n ⊆ S" by (induct n) (auto simp add: Sseq_def) qed have Sseq_pick: "!!n. pick n ∈ Sseq n" proof - fix n show "pick n ∈ Sseq n" proof (unfold pick_def, rule someI_ex) from Sseq_inf have "infinite (Sseq n)" . then have "Sseq n ≠ {}" by auto then show "∃x. x ∈ Sseq n" by auto qed qed with Sseq_S have rng: "range pick ⊆ S" by auto have pick_Sseq_gt: "!!n m. pick n ∉ Sseq (n + Suc m)" proof - fix n m show "pick n ∉ Sseq (n + Suc m)" by (induct m) (auto simp add: Sseq_def pick_def) qed have pick_pick: "!!n m. pick n ≠ pick (n + Suc m)" proof - fix n m from Sseq_pick have "pick (n + Suc m) ∈ Sseq (n + Suc m)" . moreover from pick_Sseq_gt have "pick n ∉ Sseq (n + Suc m)" . ultimately show "pick n ≠ pick (n + Suc m)" by auto qed have inj: "inj pick" proof (rule linorder_injI) fix i j :: nat assume "i < j" show "pick i ≠ pick j" proof assume eq: "pick i = pick j" from `i < j` obtain k where "j = i + Suc k" by (auto simp add: less_iff_Suc_add) with pick_pick have "pick i ≠ pick j" by simp with eq show False by simp qed qed from rng inj show ?thesis by auto qed lemma infinite_iff_countable_subset: "infinite S = (∃f. inj (f::nat => 'a) ∧ range f ⊆ S)" by (auto simp add: infinite_countable_subset range_inj_infinite infinite_super) text {* For any function with infinite domain and finite range there is some element that is the image of infinitely many domain elements. In particular, any infinite sequence of elements from a finite set contains some element that occurs infinitely often. *} lemma inf_img_fin_dom: assumes img: "finite (f`A)" and dom: "infinite A" shows "∃y ∈ f`A. infinite (f -` {y})" proof (rule ccontr) assume "¬ ?thesis" with img have "finite (UN y:f`A. f -` {y})" by (blast intro: finite_UN_I) moreover have "A ⊆ (UN y:f`A. f -` {y})" by auto moreover note dom ultimately show False by (simp add: infinite_super) qed lemma inf_img_fin_domE: assumes "finite (f`A)" and "infinite A" obtains y where "y ∈ f`A" and "infinite (f -` {y})" using assms by (blast dest: inf_img_fin_dom) subsection "Infinitely Many and Almost All" text {* We often need to reason about the existence of infinitely many (resp., all but finitely many) objects satisfying some predicate, so we introduce corresponding binders and their proof rules. *} definition Inf_many :: "('a => bool) => bool" (binder "INFM " 10) where "Inf_many P = infinite {x. P x}" definition Alm_all :: "('a => bool) => bool" (binder "MOST " 10) where "Alm_all P = (¬ (INFM x. ¬ P x))" notation (xsymbols) Inf_many (binder "∃∞" 10) and Alm_all (binder "∀∞" 10) notation (HTML output) Inf_many (binder "∃∞" 10) and Alm_all (binder "∀∞" 10) lemma INF_EX: "(∃∞x. P x) ==> (∃x. P x)" unfolding Inf_many_def proof (rule ccontr) assume inf: "infinite {x. P x}" assume "¬ ?thesis" then have "{x. P x} = {}" by simp then have "finite {x. P x}" by simp with inf show False by simp qed lemma MOST_iff_finiteNeg: "(∀∞x. P x) = finite {x. ¬ P x}" by (simp add: Alm_all_def Inf_many_def) lemma ALL_MOST: "∀x. P x ==> ∀∞x. P x" by (simp add: MOST_iff_finiteNeg) lemma INF_mono: assumes inf: "∃∞x. P x" and q: "!!x. P x ==> Q x" shows "∃∞x. Q x" proof - from inf have "infinite {x. P x}" unfolding Inf_many_def . moreover from q have "{x. P x} ⊆ {x. Q x}" by auto ultimately show ?thesis by (simp add: Inf_many_def infinite_super) qed lemma MOST_mono: "∀∞x. P x ==> (!!x. P x ==> Q x) ==> ∀∞x. Q x" unfolding Alm_all_def by (blast intro: INF_mono) lemma INF_nat: "(∃∞n. P (n::nat)) = (∀m. ∃n. m<n ∧ P n)" by (simp add: Inf_many_def infinite_nat_iff_unbounded) lemma INF_nat_le: "(∃∞n. P (n::nat)) = (∀m. ∃n. m≤n ∧ P n)" by (simp add: Inf_many_def infinite_nat_iff_unbounded_le) lemma MOST_nat: "(∀∞n. P (n::nat)) = (∃m. ∀n. m<n --> P n)" by (simp add: Alm_all_def INF_nat) lemma MOST_nat_le: "(∀∞n. P (n::nat)) = (∃m. ∀n. m≤n --> P n)" by (simp add: Alm_all_def INF_nat_le) subsection "Enumeration of an Infinite Set" text {* The set's element type must be wellordered (e.g. the natural numbers). *} consts enumerate :: "'a::wellorder set => (nat => 'a::wellorder)" primrec 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 (fastsimp intro: LeastI dest!: infinite_imp_nonempty) apply (fastsimp iff: finite_Diff_singleton) done declare enumerate_0 [simp del] enumerate_Suc [simp del] lemma enumerate_step: "infinite S ==> enumerate S n < enumerate S (Suc n)" apply (induct n arbitrary: S) apply (rule order_le_neq_trans) apply (simp add: enumerate_0 Least_le enumerate_in_set) apply (simp only: enumerate_Suc') apply (subgoal_tac "enumerate (S - {enumerate S 0}) 0 : S - {enumerate S 0}") apply (blast intro: sym) apply (simp add: enumerate_in_set del: Diff_iff) apply (simp add: enumerate_Suc') 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 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
lemma infinite_imp_nonempty:
infinite S ==> S ≠ {}
lemma infinite_remove:
infinite S ==> infinite (S - {a})
lemma Diff_infinite_finite:
[| finite T; infinite S |] ==> infinite (S - T)
lemma Un_infinite:
infinite S ==> infinite (S ∪ T)
lemma infinite_super:
[| S ⊆ T; infinite S |] ==> infinite T
lemma finite_nat_bounded:
finite S ==> ∃k. S ⊆ {..<k}
lemma finite_nat_iff_bounded:
finite S = (∃k. S ⊆ {..<k})
lemma finite_nat_iff_bounded_le:
finite S = (∃k. S ⊆ {..k})
lemma infinite_nat_iff_unbounded:
infinite S = (∀m. ∃n>m. n ∈ S)
lemma infinite_nat_iff_unbounded_le:
infinite S = (∀m. ∃n≥m. n ∈ S)
lemma unbounded_k_infinite:
∀m>k. ∃n>m. n ∈ S ==> infinite S
lemma nat_infinite:
infinite UNIV
lemma nat_not_finite:
finite UNIV ==> R
lemma range_inj_infinite:
inj f ==> infinite (range f)
lemma int_infinite:
infinite UNIV
lemma linorder_injI:
(!!x y. x < y ==> f x ≠ f y) ==> inj f
lemma infinite_countable_subset:
infinite S ==> ∃f. inj f ∧ range f ⊆ S
lemma infinite_iff_countable_subset:
infinite S = (∃f. inj f ∧ range f ⊆ S)
lemma inf_img_fin_dom:
[| finite (f ` A); infinite A |] ==> ∃y∈f ` A. infinite (f -` {y})
lemma inf_img_fin_domE:
[| finite (f ` A); infinite A;
!!y. [| y ∈ f ` A; infinite (f -` {y}) |] ==> thesis |]
==> thesis
lemma INF_EX:
∃∞x. P x ==> ∃x. P x
lemma MOST_iff_finiteNeg:
(∀∞x. P x) = finite {x. ¬ P x}
lemma ALL_MOST:
∀x. P x ==> ∀∞x. P x
lemma INF_mono:
[| ∃∞x. P x; !!x. P x ==> Q x |] ==> ∃∞x. Q x
lemma MOST_mono:
[| ∀∞x. P x; !!x. P x ==> Q x |] ==> ∀∞x. Q x
lemma INF_nat:
(∃∞n. P n) = (∀m. ∃n>m. P n)
lemma INF_nat_le:
(∃∞n. P n) = (∀m. ∃n≥m. P n)
lemma MOST_nat:
(∀∞n. P n) = (∃m. ∀n>m. P n)
lemma MOST_nat_le:
(∀∞n. P n) = (∃m. ∀n≥m. P n)
lemma enumerate_Suc':
enumerate S (Suc n) = enumerate (S - {enumerate S 0}) n
lemma enumerate_in_set:
infinite S ==> enumerate S n ∈ S
lemma enumerate_step:
infinite S ==> enumerate S n < enumerate S (Suc n)
lemma enumerate_mono:
[| m < n; infinite S |] ==> enumerate S m < enumerate S n
lemma atmost_one_empty:
S = {} ==> atmost_one S
lemma atmost_one_singleton:
S = {x} ==> atmost_one S
lemma atmost_one_unique:
[| atmost_one S; x ∈ S; y ∈ S |] ==> y = x