paulson@14442: (* Title: HOL/Infnite_Set.thy paulson@14442: ID: $Id$ wenzelm@14896: Author: Stephan Merz paulson@14442: *) paulson@14442: paulson@14442: header {* Infnite Sets and Related Concepts*} paulson@14442: nipkow@15131: theory Infinite_Set nipkow@15131: import Hilbert_Choice Finite_Set SetInterval nipkow@15131: begin paulson@14442: paulson@14442: subsection "Infinite Sets" paulson@14442: paulson@14442: text {* Some elementary facts about infinite sets, by Stefan Merz. *} paulson@14442: paulson@14442: syntax paulson@14442: infinite :: "'a set \ bool" paulson@14442: translations paulson@14442: "infinite S" == "S \ Finites" paulson@14442: paulson@14442: text {* paulson@14442: Infinite sets are non-empty, and if we remove some elements paulson@14442: from an infinite set, the result is still infinite. paulson@14442: *} paulson@14442: paulson@14442: lemma infinite_nonempty: paulson@14442: "\ (infinite {})" paulson@14442: by simp paulson@14442: paulson@14442: lemma infinite_remove: paulson@14442: "infinite S \ infinite (S - {a})" paulson@14442: by simp paulson@14442: paulson@14442: lemma Diff_infinite_finite: paulson@14442: assumes T: "finite T" and S: "infinite S" paulson@14442: shows "infinite (S-T)" paulson@14442: using T paulson@14442: proof (induct) paulson@14442: from S paulson@14442: show "infinite (S - {})" by auto paulson@14442: next paulson@14442: fix T x paulson@14442: assume ih: "infinite (S-T)" paulson@14442: have "S - (insert x T) = (S-T) - {x}" paulson@14442: by (rule Diff_insert) paulson@14442: with ih paulson@14442: show "infinite (S - (insert x T))" paulson@14442: by (simp add: infinite_remove) paulson@14442: qed paulson@14442: paulson@14442: lemma Un_infinite: paulson@14442: "infinite S \ infinite (S \ T)" paulson@14442: by simp paulson@14442: paulson@14442: lemma infinite_super: paulson@14442: assumes T: "S \ T" and S: "infinite S" paulson@14442: shows "infinite T" paulson@14442: proof (rule ccontr) paulson@14442: assume "\(infinite T)" paulson@14442: with T paulson@14442: have "finite S" by (simp add: finite_subset) paulson@14442: with S paulson@14442: show False by simp paulson@14442: qed paulson@14442: paulson@14442: text {* paulson@14442: As a concrete example, we prove that the set of natural paulson@14442: numbers is infinite. paulson@14442: *} paulson@14442: paulson@14442: lemma finite_nat_bounded: paulson@14442: assumes S: "finite (S::nat set)" nipkow@15045: shows "\k. S \ {..k. ?bounded S k") paulson@14442: using S paulson@14442: proof (induct) paulson@14442: have "?bounded {} 0" by simp paulson@14442: thus "\k. ?bounded {} k" .. paulson@14442: next paulson@14442: fix S x paulson@14442: assume "\k. ?bounded S k" paulson@14442: then obtain k where k: "?bounded S k" .. paulson@14442: show "\k. ?bounded (insert x S) k" paulson@14442: proof (cases "xk. S \ {.. {..k. S \ {..k})" (is "?lhs = ?rhs") paulson@14442: proof paulson@14442: assume ?lhs nipkow@15045: then obtain k where "S \ {.. {..k}" by auto paulson@14442: thus ?rhs .. paulson@14442: next paulson@14442: assume ?rhs paulson@14442: then obtain k where "S \ {..k}" .. paulson@14442: thus "finite S" paulson@14442: by (rule finite_subset, simp) paulson@14442: qed paulson@14442: paulson@14442: lemma infinite_nat_iff_unbounded: paulson@14442: "infinite (S::nat set) = (\m. \n. m n\S)" paulson@14442: (is "?lhs = ?rhs") paulson@14442: proof paulson@14442: assume inf: ?lhs paulson@14442: show ?rhs paulson@14442: proof (rule ccontr) paulson@14442: assume "\ ?rhs" paulson@14442: then obtain m where m: "\n. m n\S" by blast paulson@14442: hence "S \ {..m}" paulson@14442: by (auto simp add: sym[OF not_less_iff_le]) paulson@14442: with inf show "False" paulson@14442: by (simp add: finite_nat_iff_bounded_le) paulson@14442: qed paulson@14442: next paulson@14442: assume unbounded: ?rhs paulson@14442: show ?lhs paulson@14442: proof paulson@14442: assume "finite S" paulson@14442: then obtain m where "S \ {..m}" paulson@14442: by (auto simp add: finite_nat_iff_bounded_le) paulson@14442: hence "\n. m n\S" by auto paulson@14442: with unbounded show "False" by blast paulson@14442: qed paulson@14442: qed paulson@14442: paulson@14442: lemma infinite_nat_iff_unbounded_le: paulson@14442: "infinite (S::nat set) = (\m. \n. m\n \ n\S)" paulson@14442: (is "?lhs = ?rhs") paulson@14442: proof paulson@14442: assume inf: ?lhs paulson@14442: show ?rhs paulson@14442: proof paulson@14442: fix m paulson@14442: from inf obtain n where "m n\S" paulson@14442: by (auto simp add: infinite_nat_iff_unbounded) paulson@14442: hence "m\n \ n\S" by auto paulson@14442: thus "\n. m \ n \ n \ S" .. paulson@14442: qed paulson@14442: next paulson@14442: assume unbounded: ?rhs paulson@14442: show ?lhs paulson@14442: proof (auto simp add: infinite_nat_iff_unbounded) paulson@14442: fix m paulson@14442: from unbounded obtain n where "(Suc m)\n \ n\S" paulson@14442: by blast paulson@14442: hence "m n\S" by auto paulson@14442: thus "\n. m < n \ n \ S" .. paulson@14442: qed paulson@14442: qed paulson@14442: paulson@14442: text {* paulson@14442: For a set of natural numbers to be infinite, it is enough wenzelm@14957: to know that for any number larger than some @{text k}, there paulson@14442: is some larger number that is an element of the set. paulson@14442: *} paulson@14442: paulson@14442: lemma unbounded_k_infinite: paulson@14442: assumes k: "\m. k (\n. m n\S)" paulson@14442: shows "infinite (S::nat set)" paulson@14442: proof (auto simp add: infinite_nat_iff_unbounded) paulson@14442: fix m show "\n. m n\S" paulson@14442: proof (cases "k n\S" by auto paulson@14442: with False have "m n\S" by auto paulson@14442: thus ?thesis .. paulson@14442: qed paulson@14442: qed paulson@14442: paulson@14442: theorem nat_infinite [simp]: paulson@14442: "infinite (UNIV :: nat set)" paulson@14442: by (auto simp add: infinite_nat_iff_unbounded) paulson@14442: paulson@14442: theorem nat_not_finite [elim]: paulson@14442: "finite (UNIV::nat set) \ R" paulson@14442: by simp paulson@14442: paulson@14442: text {* paulson@14442: Every infinite set contains a countable subset. More precisely wenzelm@14957: we show that a set @{text S} is infinite if and only if there exists wenzelm@14957: an injective function from the naturals into @{text S}. paulson@14442: *} paulson@14442: paulson@14442: lemma range_inj_infinite: paulson@14442: "inj (f::nat \ 'a) \ infinite (range f)" paulson@14442: proof paulson@14442: assume "inj f" paulson@14442: and "finite (range f)" paulson@14442: hence "finite (UNIV::nat set)" paulson@14442: by (auto intro: finite_imageD simp del: nat_infinite) paulson@14442: thus "False" by simp paulson@14442: qed paulson@14442: paulson@14442: text {* paulson@14442: The ``only if'' direction is harder because it requires the paulson@14442: construction of a sequence of pairwise different elements of wenzelm@14957: an infinite set @{text S}. The idea is to construct a sequence of wenzelm@14957: non-empty and infinite subsets of @{text S} obtained by successively wenzelm@14957: removing elements of @{text S}. paulson@14442: *} paulson@14442: paulson@14442: lemma linorder_injI: paulson@14442: assumes hyp: "\x y. x < (y::'a::linorder) \ f x \ f y" paulson@14442: shows "inj f" paulson@14442: proof (rule inj_onI) paulson@14442: fix x y paulson@14442: assume f_eq: "f x = f y" paulson@14442: show "x = y" paulson@14442: proof (rule linorder_cases) paulson@14442: assume "x < y" paulson@14442: with hyp have "f x \ f y" by blast paulson@14442: with f_eq show ?thesis by simp paulson@14442: next paulson@14442: assume "x = y" paulson@14442: thus ?thesis . paulson@14442: next paulson@14442: assume "y < x" paulson@14442: with hyp have "f y \ f x" by blast paulson@14442: with f_eq show ?thesis by simp paulson@14442: qed paulson@14442: qed paulson@14442: paulson@14442: lemma infinite_countable_subset: paulson@14442: assumes inf: "infinite (S::'a set)" paulson@14442: shows "\f. inj (f::nat \ 'a) \ range f \ S" paulson@14442: proof - wenzelm@14766: def Sseq \ "nat_rec S (\n T. T - {SOME e. e \ T})" wenzelm@14766: def pick \ "\n. (SOME e. e \ Sseq n)" paulson@14442: have Sseq_inf: "\n. infinite (Sseq n)" paulson@14442: proof - paulson@14442: fix n paulson@14442: show "infinite (Sseq n)" paulson@14442: proof (induct n) paulson@14442: from inf show "infinite (Sseq 0)" paulson@14442: by (simp add: Sseq_def) paulson@14442: next paulson@14442: fix n paulson@14442: assume "infinite (Sseq n)" thus "infinite (Sseq (Suc n))" paulson@14442: by (simp add: Sseq_def infinite_remove) paulson@14442: qed paulson@14442: qed paulson@14442: have Sseq_S: "\n. Sseq n \ S" paulson@14442: proof - paulson@14442: fix n paulson@14442: show "Sseq n \ S" paulson@14442: by (induct n, auto simp add: Sseq_def) paulson@14442: qed paulson@14442: have Sseq_pick: "\n. pick n \ Sseq n" paulson@14442: proof - paulson@14442: fix n paulson@14442: show "pick n \ Sseq n" paulson@14442: proof (unfold pick_def, rule someI_ex) paulson@14442: from Sseq_inf have "infinite (Sseq n)" . paulson@14442: hence "Sseq n \ {}" by auto paulson@14442: thus "\x. x \ Sseq n" by auto paulson@14442: qed paulson@14442: qed paulson@14442: with Sseq_S have rng: "range pick \ S" paulson@14442: by auto paulson@14442: have pick_Sseq_gt: "\n m. pick n \ Sseq (n + Suc m)" paulson@14442: proof - paulson@14442: fix n m paulson@14442: show "pick n \ Sseq (n + Suc m)" paulson@14442: by (induct m, auto simp add: Sseq_def pick_def) paulson@14442: qed paulson@14442: have pick_pick: "\n m. pick n \ pick (n + Suc m)" paulson@14442: proof - paulson@14442: fix n m paulson@14442: from Sseq_pick have "pick (n + Suc m) \ Sseq (n + Suc m)" . paulson@14442: moreover from pick_Sseq_gt paulson@14442: have "pick n \ Sseq (n + Suc m)" . paulson@14442: ultimately show "pick n \ pick (n + Suc m)" paulson@14442: by auto paulson@14442: qed paulson@14442: have inj: "inj pick" paulson@14442: proof (rule linorder_injI) paulson@14442: show "\i j. i<(j::nat) \ pick i \ pick j" paulson@14442: proof (clarify) paulson@14442: fix i j paulson@14442: assume ij: "i<(j::nat)" paulson@14442: and eq: "pick i = pick j" paulson@14442: from ij obtain k where "j = i + (Suc k)" paulson@14442: by (auto simp add: less_iff_Suc_add) paulson@14442: with pick_pick have "pick i \ pick j" by simp paulson@14442: with eq show "False" by simp paulson@14442: qed paulson@14442: qed paulson@14442: from rng inj show ?thesis by auto paulson@14442: qed paulson@14442: paulson@14442: theorem infinite_iff_countable_subset: paulson@14442: "infinite S = (\f. inj (f::nat \ 'a) \ range f \ S)" paulson@14442: (is "?lhs = ?rhs") paulson@14442: by (auto simp add: infinite_countable_subset paulson@14442: range_inj_infinite infinite_super) paulson@14442: paulson@14442: text {* paulson@14442: For any function with infinite domain and finite range paulson@14442: there is some element that is the image of infinitely paulson@14442: many domain elements. In particular, any infinite sequence paulson@14442: of elements from a finite set contains some element that paulson@14442: occurs infinitely often. paulson@14442: *} paulson@14442: paulson@14442: theorem inf_img_fin_dom: paulson@14442: assumes img: "finite (f`A)" and dom: "infinite A" paulson@14442: shows "\y \ f`A. infinite (f -` {y})" paulson@14442: proof (rule ccontr) paulson@14442: assume "\ (\y\f ` A. infinite (f -` {y}))" paulson@14442: with img have "finite (UN y:f`A. f -` {y})" paulson@14442: by (blast intro: finite_UN_I) paulson@14442: moreover have "A \ (UN y:f`A. f -` {y})" by auto paulson@14442: moreover note dom paulson@14442: ultimately show "False" paulson@14442: by (simp add: infinite_super) paulson@14442: qed paulson@14442: paulson@14442: theorems inf_img_fin_domE = inf_img_fin_dom[THEN bexE] paulson@14442: paulson@14442: paulson@14442: subsection "Infinitely Many and Almost All" paulson@14442: paulson@14442: text {* paulson@14442: We often need to reason about the existence of infinitely many paulson@14442: (resp., all but finitely many) objects satisfying some predicate, paulson@14442: so we introduce corresponding binders and their proof rules. paulson@14442: *} paulson@14442: paulson@14442: consts paulson@14442: Inf_many :: "('a \ bool) \ bool" (binder "INF " 10) paulson@14442: Alm_all :: "('a \ bool) \ bool" (binder "MOST " 10) paulson@14442: paulson@14442: defs paulson@14442: INF_def: "Inf_many P \ infinite {x. P x}" paulson@14442: MOST_def: "Alm_all P \ \(INF x. \ P x)" paulson@14442: paulson@14442: syntax (xsymbols) paulson@14442: "MOST " :: "[idts, bool] \ bool" ("(3\\<^sub>\_./ _)" [0,10] 10) paulson@14442: "INF " :: "[idts, bool] \ bool" ("(3\\<^sub>\_./ _)" [0,10] 10) paulson@14442: kleing@14565: syntax (HTML output) kleing@14565: "MOST " :: "[idts, bool] \ bool" ("(3\\<^sub>\_./ _)" [0,10] 10) kleing@14565: "INF " :: "[idts, bool] \ bool" ("(3\\<^sub>\_./ _)" [0,10] 10) kleing@14565: paulson@14442: lemma INF_EX: paulson@14442: "(\\<^sub>\x. P x) \ (\x. P x)" paulson@14442: proof (unfold INF_def, rule ccontr) paulson@14442: assume inf: "infinite {x. P x}" paulson@14442: and notP: "\(\x. P x)" paulson@14442: from notP have "{x. P x} = {}" by simp paulson@14442: hence "finite {x. P x}" by simp paulson@14442: with inf show "False" by simp paulson@14442: qed paulson@14442: paulson@14442: lemma MOST_iff_finiteNeg: paulson@14442: "(\\<^sub>\x. P x) = finite {x. \ P x}" paulson@14442: by (simp add: MOST_def INF_def) paulson@14442: paulson@14442: lemma ALL_MOST: paulson@14442: "\x. P x \ \\<^sub>\x. P x" paulson@14442: by (simp add: MOST_iff_finiteNeg) paulson@14442: paulson@14442: lemma INF_mono: paulson@14442: assumes inf: "\\<^sub>\x. P x" and q: "\x. P x \ Q x" paulson@14442: shows "\\<^sub>\x. Q x" paulson@14442: proof - paulson@14442: from inf have "infinite {x. P x}" by (unfold INF_def) paulson@14442: moreover from q have "{x. P x} \ {x. Q x}" by auto paulson@14442: ultimately show ?thesis paulson@14442: by (simp add: INF_def infinite_super) paulson@14442: qed paulson@14442: paulson@14442: lemma MOST_mono: paulson@14442: "\ \\<^sub>\x. P x; \x. P x \ Q x \ \ \\<^sub>\x. Q x" paulson@14442: by (unfold MOST_def, blast intro: INF_mono) paulson@14442: paulson@14442: lemma INF_nat: "(\\<^sub>\n. P (n::nat)) = (\m. \n. m P n)" paulson@14442: by (simp add: INF_def infinite_nat_iff_unbounded) paulson@14442: paulson@14442: lemma INF_nat_le: "(\\<^sub>\n. P (n::nat)) = (\m. \n. m\n \ P n)" paulson@14442: by (simp add: INF_def infinite_nat_iff_unbounded_le) paulson@14442: paulson@14442: lemma MOST_nat: "(\\<^sub>\n. P (n::nat)) = (\m. \n. m P n)" paulson@14442: by (simp add: MOST_def INF_nat) paulson@14442: paulson@14442: lemma MOST_nat_le: "(\\<^sub>\n. P (n::nat)) = (\m. \n. m\n \ P n)" paulson@14442: by (simp add: MOST_def INF_nat_le) paulson@14442: paulson@14442: paulson@14442: subsection "Miscellaneous" paulson@14442: paulson@14442: text {* paulson@14442: A few trivial lemmas about sets that contain at most one element. paulson@14442: These simplify the reasoning about deterministic automata. paulson@14442: *} paulson@14442: paulson@14442: constdefs paulson@14442: atmost_one :: "'a set \ bool" paulson@14442: "atmost_one S \ \x y. x\S \ y\S \ x=y" paulson@14442: paulson@14442: lemma atmost_one_empty: "S={} \ atmost_one S" paulson@14442: by (simp add: atmost_one_def) paulson@14442: paulson@14442: lemma atmost_one_singleton: "S = {x} \ atmost_one S" paulson@14442: by (simp add: atmost_one_def) paulson@14442: paulson@14442: lemma atmost_one_unique [elim]: "\ atmost_one S; x \ S; y \ S \ \ y=x" paulson@14442: by (simp add: atmost_one_def) paulson@14442: paulson@14442: end