# HG changeset patch # User paulson # Date 1078740718 -3600 # Node ID 04135b0c06ffe4b89920bbd749c20ac1f72a6a2f # Parent 40d7ae9563fda6ace948ad775090de0a57fdda63 new theory of infinite sets diff -r 40d7ae9563fd -r 04135b0c06ff src/HOL/Infinite_Set.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Infinite_Set.thy Mon Mar 08 11:11:58 2004 +0100 @@ -0,0 +1,427 @@ +(* Title: HOL/Infnite_Set.thy + ID: $Id$ + Author: Stefan Merz +*) + +header {* Infnite Sets and Related Concepts*} + +theory Infinite_Set = Hilbert_Choice + Finite_Set: + +subsection "Infinite Sets" + +text {* Some elementary facts about infinite sets, by Stefan Merz. *} + +syntax + infinite :: "'a set \ bool" +translations + "infinite S" == "S \ Finites" + +text {* + Infinite sets are non-empty, and if we remove some elements + from an infinite set, the result is still infinite. +*} + +lemma infinite_nonempty: + "\ (infinite {})" +by simp + +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 (rule ccontr) + assume "\(infinite 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 + thus "\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 "xk. S \ {..k(})" (is "?lhs = ?rhs") +proof + assume ?lhs + thus ?rhs by (rule finite_nat_bounded) +next + assume ?rhs + then obtain k where "S \ {..k(}" .. + thus "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) + hence "S \ {..k}" by auto + thus ?rhs .. +next + assume ?rhs + then obtain k where "S \ {..k}" .. + thus "finite S" + by (rule finite_subset, simp) +qed + +lemma infinite_nat_iff_unbounded: + "infinite (S::nat set) = (\m. \n. m n\S)" + (is "?lhs = ?rhs") +proof + assume inf: ?lhs + show ?rhs + proof (rule ccontr) + assume "\ ?rhs" + then obtain m where m: "\n. m n\S" by blast + hence "S \ {..m}" + by (auto simp add: sym[OF not_less_iff_le]) + with inf show "False" + by (simp add: finite_nat_iff_bounded_le) + qed +next + assume unbounded: ?rhs + show ?lhs + proof + assume "finite S" + then obtain m where "S \ {..m}" + by (auto simp add: finite_nat_iff_bounded_le) + hence "\n. m n\S" by auto + with unbounded 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 inf: ?lhs + show ?rhs + proof + fix m + from inf obtain n where "m n\S" + by (auto simp add: infinite_nat_iff_unbounded) + hence "m\n \ n\S" by auto + thus "\n. m \ n \ n \ S" .. + qed +next + assume unbounded: ?rhs + show ?lhs + proof (auto simp add: infinite_nat_iff_unbounded) + fix m + from unbounded obtain n where "(Suc m)\n \ n\S" + by blast + hence "m n\S" by auto + thus "\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 $k$, there + is some larger number that is an element of the set. +*} + +lemma unbounded_k_infinite: + assumes k: "\m. k (\n. m n\S)" + shows "infinite (S::nat set)" +proof (auto simp add: infinite_nat_iff_unbounded) + fix m show "\n. m n\S" + proof (cases "k n\S" by auto + with False have "m n\S" by auto + thus ?thesis .. + qed +qed + +theorem nat_infinite [simp]: + "infinite (UNIV :: nat set)" +by (auto simp add: infinite_nat_iff_unbounded) + +theorem 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 $S$ is infinite if and only if there exists + an injective function from the naturals into $S$. +*} + +lemma range_inj_infinite: + "inj (f::nat \ 'a) \ infinite (range f)" +proof + assume "inj f" + and "finite (range f)" + hence "finite (UNIV::nat set)" + by (auto intro: finite_imageD simp del: nat_infinite) + thus "False" by simp +qed + +text {* + The ``only if'' direction is harder because it requires the + construction of a sequence of pairwise different elements of + an infinite set $S$. The idea is to construct a sequence of + non-empty and infinite subsets of $S$ obtained by successively + removing elements of $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" + thus ?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 - {\ e. e \ T})" + def pick \ "\n. (\ 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)" thus "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)" . + hence "Sseq n \ {}" by auto + thus "\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) + show "\i j. i<(j::nat) \ pick i \ pick j" + proof (clarify) + fix i j + assume ij: "i<(j::nat)" + and eq: "pick i = pick j" + from ij 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 + +theorem infinite_iff_countable_subset: + "infinite S = (\f. inj (f::nat \ 'a) \ range f \ S)" + (is "?lhs = ?rhs") +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. +*} + +theorem inf_img_fin_dom: + assumes img: "finite (f`A)" and dom: "infinite A" + shows "\y \ f`A. infinite (f -` {y})" +proof (rule ccontr) + assume "\ (\y\f ` A. infinite (f -` {y}))" + 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 + +theorems inf_img_fin_domE = inf_img_fin_dom[THEN bexE] + + +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. +*} + +consts + Inf_many :: "('a \ bool) \ bool" (binder "INF " 10) + Alm_all :: "('a \ bool) \ bool" (binder "MOST " 10) + +defs + INF_def: "Inf_many P \ infinite {x. P x}" + MOST_def: "Alm_all P \ \(INF x. \ P x)" + +syntax (xsymbols) + "MOST " :: "[idts, bool] \ bool" ("(3\\<^sub>\_./ _)" [0,10] 10) + "INF " :: "[idts, bool] \ bool" ("(3\\<^sub>\_./ _)" [0,10] 10) + +lemma INF_EX: + "(\\<^sub>\x. P x) \ (\x. P x)" +proof (unfold INF_def, rule ccontr) + assume inf: "infinite {x. P x}" + and notP: "\(\x. P x)" + from notP have "{x. P x} = {}" by simp + hence "finite {x. P x}" by simp + with inf show "False" by simp +qed + +lemma MOST_iff_finiteNeg: + "(\\<^sub>\x. P x) = finite {x. \ P x}" +by (simp add: MOST_def INF_def) + +lemma ALL_MOST: + "\x. P x \ \\<^sub>\x. P x" +by (simp add: MOST_iff_finiteNeg) + +lemma INF_mono: + assumes inf: "\\<^sub>\x. P x" and q: "\x. P x \ Q x" + shows "\\<^sub>\x. Q x" +proof - + from inf have "infinite {x. P x}" by (unfold INF_def) + moreover from q have "{x. P x} \ {x. Q x}" by auto + ultimately show ?thesis + by (simp add: INF_def infinite_super) +qed + +lemma MOST_mono: + "\ \\<^sub>\x. P x; \x. P x \ Q x \ \ \\<^sub>\x. Q x" +by (unfold MOST_def, blast intro: INF_mono) + +lemma INF_nat: "(\\<^sub>\n. P (n::nat)) = (\m. \n. m P n)" +by (simp add: INF_def infinite_nat_iff_unbounded) + +lemma INF_nat_le: "(\\<^sub>\n. P (n::nat)) = (\m. \n. m\n \ P n)" +by (simp add: INF_def infinite_nat_iff_unbounded_le) + +lemma MOST_nat: "(\\<^sub>\n. P (n::nat)) = (\m. \n. m P n)" +by (simp add: MOST_def INF_nat) + +lemma MOST_nat_le: "(\\<^sub>\n. P (n::nat)) = (\m. \n. m\n \ P n)" +by (simp add: MOST_def INF_nat_le) + + +subsection "Miscellaneous" + +text {* + A few trivial lemmas about sets that contain at most one element. + These simplify the reasoning about deterministic automata. +*} + +constdefs + atmost_one :: "'a set \ bool" + "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 diff -r 40d7ae9563fd -r 04135b0c06ff src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Sat Mar 06 19:32:21 2004 +0100 +++ b/src/HOL/IsaMakefile Mon Mar 08 11:11:58 2004 +0100 @@ -84,7 +84,7 @@ Divides.thy Extraction.thy Finite_Set.ML Finite_Set.thy \ Fun.thy Gfp.ML Gfp.thy \ Hilbert_Choice.thy Hilbert_Choice_lemmas.ML HOL.ML \ - HOL.thy HOL_lemmas.ML Inductive.thy Integ/Bin.thy \ + HOL.thy HOL_lemmas.ML Inductive.thy Infinite_Set.thy Integ/Bin.thy \ Integ/cooper_dec.ML Integ/cooper_proof.ML \ Integ/Equiv.thy Integ/IntArith.thy Integ/IntDef.thy \ Integ/IntDiv.thy Integ/NatBin.thy Integ/NatSimprocs.thy Integ/Parity.thy \