author | immler |
Tue, 15 May 2018 11:33:43 +0200 | |
changeset 68188 | 2af1f142f855 |
parent 67408 | 4a4c14b24800 |
child 68406 | 6beb45f6cf67 |
permissions | -rw-r--r-- |
27407
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
huffman
parents:
27368
diff
changeset
|
1 |
(* Title: HOL/Library/Infinite_Set.thy |
20809 | 2 |
Author: Stephan Merz |
3 |
*) |
|
4 |
||
60500 | 5 |
section \<open>Infinite Sets and Related Concepts\<close> |
20809 | 6 |
|
7 |
theory Infinite_Set |
|
64967 | 8 |
imports Main |
20809 | 9 |
begin |
10 |
||
64967 | 11 |
subsection \<open>The set of natural numbers is infinite\<close> |
20809 | 12 |
|
64967 | 13 |
lemma infinite_nat_iff_unbounded_le: "infinite S \<longleftrightarrow> (\<forall>m. \<exists>n\<ge>m. n \<in> S)" |
14 |
for S :: "nat set" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
15 |
using frequently_cofinite[of "\<lambda>x. x \<in> S"] |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
16 |
by (simp add: cofinite_eq_sequentially frequently_def eventually_sequentially) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
17 |
|
64967 | 18 |
lemma infinite_nat_iff_unbounded: "infinite S \<longleftrightarrow> (\<forall>m. \<exists>n>m. n \<in> S)" |
19 |
for S :: "nat set" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
20 |
using frequently_cofinite[of "\<lambda>x. x \<in> S"] |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
21 |
by (simp add: cofinite_eq_sequentially frequently_def eventually_at_top_dense) |
20809 | 22 |
|
64967 | 23 |
lemma finite_nat_iff_bounded: "finite S \<longleftrightarrow> (\<exists>k. S \<subseteq> {..<k})" |
24 |
for S :: "nat set" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
25 |
using infinite_nat_iff_unbounded_le[of S] by (simp add: subset_eq) (metis not_le) |
20809 | 26 |
|
64967 | 27 |
lemma finite_nat_iff_bounded_le: "finite S \<longleftrightarrow> (\<exists>k. S \<subseteq> {.. k})" |
28 |
for S :: "nat set" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
29 |
using infinite_nat_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) |
20809 | 30 |
|
64967 | 31 |
lemma finite_nat_bounded: "finite S \<Longrightarrow> \<exists>k. S \<subseteq> {..<k}" |
32 |
for S :: "nat set" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
33 |
by (simp add: finite_nat_iff_bounded) |
20809 | 34 |
|
61762
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
35 |
|
60500 | 36 |
text \<open> |
20809 | 37 |
For a set of natural numbers to be infinite, it is enough to know |
61585 | 38 |
that for any number larger than some \<open>k\<close>, there is some larger |
20809 | 39 |
number that is an element of the set. |
60500 | 40 |
\<close> |
20809 | 41 |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
42 |
lemma unbounded_k_infinite: "\<forall>m>k. \<exists>n>m. n \<in> S \<Longrightarrow> infinite (S::nat set)" |
64967 | 43 |
apply (clarsimp simp add: finite_nat_set_iff_bounded) |
44 |
apply (drule_tac x="Suc (max m k)" in spec) |
|
45 |
using less_Suc_eq apply fastforce |
|
46 |
done |
|
20809 | 47 |
|
35056 | 48 |
lemma nat_not_finite: "finite (UNIV::nat set) \<Longrightarrow> R" |
20809 | 49 |
by simp |
50 |
||
51 |
lemma range_inj_infinite: |
|
64967 | 52 |
fixes f :: "nat \<Rightarrow> 'a" |
53 |
assumes "inj f" |
|
54 |
shows "infinite (range f)" |
|
20809 | 55 |
proof |
64967 | 56 |
assume "finite (range f)" |
57 |
from this assms have "finite (UNIV::nat set)" |
|
27407
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
huffman
parents:
27368
diff
changeset
|
58 |
by (rule finite_imageD) |
20809 | 59 |
then show False by simp |
60 |
qed |
|
61 |
||
64967 | 62 |
|
63 |
subsection \<open>The set of integers is also infinite\<close> |
|
61762
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
64 |
|
64967 | 65 |
lemma infinite_int_iff_infinite_nat_abs: "infinite S \<longleftrightarrow> infinite ((nat \<circ> abs) ` S)" |
66 |
for S :: "int set" |
|
66837 | 67 |
proof - |
68 |
have "inj_on nat (abs ` A)" for A |
|
69 |
by (rule inj_onI) auto |
|
70 |
then show ?thesis |
|
71 |
by (auto simp add: image_comp [symmetric] dest: finite_image_absD finite_imageD) |
|
72 |
qed |
|
61762
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
73 |
|
64967 | 74 |
proposition infinite_int_iff_unbounded_le: "infinite S \<longleftrightarrow> (\<forall>m. \<exists>n. \<bar>n\<bar> \<ge> m \<and> n \<in> S)" |
75 |
for S :: "int set" |
|
76 |
by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded_le o_def image_def) |
|
77 |
(metis abs_ge_zero nat_le_eq_zle le_nat_iff) |
|
61762
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
78 |
|
64967 | 79 |
proposition infinite_int_iff_unbounded: "infinite S \<longleftrightarrow> (\<forall>m. \<exists>n. \<bar>n\<bar> > m \<and> n \<in> S)" |
80 |
for S :: "int set" |
|
81 |
by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded o_def image_def) |
|
82 |
(metis (full_types) nat_le_iff nat_mono not_le) |
|
61762
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
83 |
|
64967 | 84 |
proposition finite_int_iff_bounded: "finite S \<longleftrightarrow> (\<exists>k. abs ` S \<subseteq> {..<k})" |
85 |
for S :: "int set" |
|
61762
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
86 |
using infinite_int_iff_unbounded_le[of S] by (simp add: subset_eq) (metis not_le) |
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
87 |
|
64967 | 88 |
proposition finite_int_iff_bounded_le: "finite S \<longleftrightarrow> (\<exists>k. abs ` S \<subseteq> {.. k})" |
89 |
for S :: "int set" |
|
61762
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
90 |
using infinite_int_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) |
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
91 |
|
64967 | 92 |
|
93 |
subsection \<open>Infinitely Many and Almost All\<close> |
|
20809 | 94 |
|
60500 | 95 |
text \<open> |
20809 | 96 |
We often need to reason about the existence of infinitely many |
97 |
(resp., all but finitely many) objects satisfying some predicate, so |
|
98 |
we introduce corresponding binders and their proof rules. |
|
60500 | 99 |
\<close> |
20809 | 100 |
|
64967 | 101 |
lemma not_INFM [simp]: "\<not> (INFM x. P x) \<longleftrightarrow> (MOST x. \<not> P x)" |
102 |
by (rule not_frequently) |
|
103 |
||
104 |
lemma not_MOST [simp]: "\<not> (MOST x. P x) \<longleftrightarrow> (INFM x. \<not> P x)" |
|
105 |
by (rule not_eventually) |
|
34112 | 106 |
|
107 |
lemma INFM_const [simp]: "(INFM x::'a. P) \<longleftrightarrow> P \<and> infinite (UNIV::'a set)" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
108 |
by (simp add: frequently_const_iff) |
34112 | 109 |
|
110 |
lemma MOST_const [simp]: "(MOST x::'a. P) \<longleftrightarrow> P \<or> finite (UNIV::'a set)" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
111 |
by (simp add: eventually_const_iff) |
20809 | 112 |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
113 |
lemma INFM_imp_distrib: "(INFM x. P x \<longrightarrow> Q x) \<longleftrightarrow> ((MOST x. P x) \<longrightarrow> (INFM x. Q x))" |
64967 | 114 |
by (rule frequently_imp_iff) |
34112 | 115 |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
116 |
lemma MOST_imp_iff: "MOST x. P x \<Longrightarrow> (MOST x. P x \<longrightarrow> Q x) \<longleftrightarrow> (MOST x. Q x)" |
61810 | 117 |
by (auto intro: eventually_rev_mp eventually_mono) |
34113 | 118 |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
119 |
lemma INFM_conjI: "INFM x. P x \<Longrightarrow> MOST x. Q x \<Longrightarrow> INFM x. P x \<and> Q x" |
61810 | 120 |
by (rule frequently_rev_mp[of P]) (auto elim: eventually_mono) |
34112 | 121 |
|
64967 | 122 |
|
60500 | 123 |
text \<open>Properties of quantifiers with injective functions.\<close> |
34112 | 124 |
|
53239 | 125 |
lemma INFM_inj: "INFM x. P (f x) \<Longrightarrow> inj f \<Longrightarrow> INFM x. P x" |
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
126 |
using finite_vimageI[of "{x. P x}" f] by (auto simp: frequently_cofinite) |
27407
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
huffman
parents:
27368
diff
changeset
|
127 |
|
53239 | 128 |
lemma MOST_inj: "MOST x. P x \<Longrightarrow> inj f \<Longrightarrow> MOST x. P (f x)" |
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
129 |
using finite_vimageI[of "{x. \<not> P x}" f] by (auto simp: eventually_cofinite) |
34112 | 130 |
|
64967 | 131 |
|
60500 | 132 |
text \<open>Properties of quantifiers with singletons.\<close> |
34112 | 133 |
|
134 |
lemma not_INFM_eq [simp]: |
|
135 |
"\<not> (INFM x. x = a)" |
|
136 |
"\<not> (INFM x. a = x)" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
137 |
unfolding frequently_cofinite by simp_all |
34112 | 138 |
|
139 |
lemma MOST_neq [simp]: |
|
140 |
"MOST x. x \<noteq> a" |
|
141 |
"MOST x. a \<noteq> x" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
142 |
unfolding eventually_cofinite by simp_all |
27407
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
huffman
parents:
27368
diff
changeset
|
143 |
|
34112 | 144 |
lemma INFM_neq [simp]: |
145 |
"(INFM x::'a. x \<noteq> a) \<longleftrightarrow> infinite (UNIV::'a set)" |
|
146 |
"(INFM x::'a. a \<noteq> x) \<longleftrightarrow> infinite (UNIV::'a set)" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
147 |
unfolding frequently_cofinite by simp_all |
34112 | 148 |
|
149 |
lemma MOST_eq [simp]: |
|
150 |
"(MOST x::'a. x = a) \<longleftrightarrow> finite (UNIV::'a set)" |
|
151 |
"(MOST x::'a. a = x) \<longleftrightarrow> finite (UNIV::'a set)" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
152 |
unfolding eventually_cofinite by simp_all |
34112 | 153 |
|
154 |
lemma MOST_eq_imp: |
|
155 |
"MOST x. x = a \<longrightarrow> P x" |
|
156 |
"MOST x. a = x \<longrightarrow> P x" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
157 |
unfolding eventually_cofinite by simp_all |
34112 | 158 |
|
64967 | 159 |
|
60500 | 160 |
text \<open>Properties of quantifiers over the naturals.\<close> |
27407
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
huffman
parents:
27368
diff
changeset
|
161 |
|
64967 | 162 |
lemma MOST_nat: "(\<forall>\<^sub>\<infinity>n. P n) \<longleftrightarrow> (\<exists>m. \<forall>n>m. P n)" |
163 |
for P :: "nat \<Rightarrow> bool" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
164 |
by (auto simp add: eventually_cofinite finite_nat_iff_bounded_le subset_eq not_le[symmetric]) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
165 |
|
64967 | 166 |
lemma MOST_nat_le: "(\<forall>\<^sub>\<infinity>n. P n) \<longleftrightarrow> (\<exists>m. \<forall>n\<ge>m. P n)" |
167 |
for P :: "nat \<Rightarrow> bool" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
168 |
by (auto simp add: eventually_cofinite finite_nat_iff_bounded subset_eq not_le[symmetric]) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
169 |
|
64967 | 170 |
lemma INFM_nat: "(\<exists>\<^sub>\<infinity>n. P n) \<longleftrightarrow> (\<forall>m. \<exists>n>m. P n)" |
171 |
for P :: "nat \<Rightarrow> bool" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
172 |
by (simp add: frequently_cofinite infinite_nat_iff_unbounded) |
20809 | 173 |
|
64967 | 174 |
lemma INFM_nat_le: "(\<exists>\<^sub>\<infinity>n. P n) \<longleftrightarrow> (\<forall>m. \<exists>n\<ge>m. P n)" |
175 |
for P :: "nat \<Rightarrow> bool" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
176 |
by (simp add: frequently_cofinite infinite_nat_iff_unbounded_le) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
177 |
|
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
178 |
lemma MOST_INFM: "infinite (UNIV::'a set) \<Longrightarrow> MOST x::'a. P x \<Longrightarrow> INFM x::'a. P x" |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
179 |
by (simp add: eventually_frequently) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
180 |
|
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
181 |
lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \<longleftrightarrow> (MOST n. P n)" |
64697 | 182 |
by (simp add: cofinite_eq_sequentially) |
20809 | 183 |
|
64967 | 184 |
lemma MOST_SucI: "MOST n. P n \<Longrightarrow> MOST n. P (Suc n)" |
185 |
and MOST_SucD: "MOST n. P (Suc n) \<Longrightarrow> MOST n. P n" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
186 |
by (simp_all add: MOST_Suc_iff) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
187 |
|
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
188 |
lemma MOST_ge_nat: "MOST n::nat. m \<le> n" |
66837 | 189 |
by (simp add: cofinite_eq_sequentially) |
20809 | 190 |
|
67408 | 191 |
\<comment> \<open>legacy names\<close> |
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
192 |
lemma Inf_many_def: "Inf_many P \<longleftrightarrow> infinite {x. P x}" by (fact frequently_cofinite) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
193 |
lemma Alm_all_def: "Alm_all P \<longleftrightarrow> \<not> (INFM x. \<not> P x)" by simp |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
194 |
lemma INFM_iff_infinite: "(INFM x. P x) \<longleftrightarrow> infinite {x. P x}" by (fact frequently_cofinite) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
195 |
lemma MOST_iff_cofinite: "(MOST x. P x) \<longleftrightarrow> finite {x. \<not> P x}" by (fact eventually_cofinite) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
196 |
lemma INFM_EX: "(\<exists>\<^sub>\<infinity>x. P x) \<Longrightarrow> (\<exists>x. P x)" by (fact frequently_ex) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
197 |
lemma ALL_MOST: "\<forall>x. P x \<Longrightarrow> \<forall>\<^sub>\<infinity>x. P x" by (fact always_eventually) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
198 |
lemma INFM_mono: "\<exists>\<^sub>\<infinity>x. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> \<exists>\<^sub>\<infinity>x. Q x" by (fact frequently_elim1) |
61810 | 199 |
lemma MOST_mono: "\<forall>\<^sub>\<infinity>x. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> \<forall>\<^sub>\<infinity>x. Q x" by (fact eventually_mono) |
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
200 |
lemma INFM_disj_distrib: "(\<exists>\<^sub>\<infinity>x. P x \<or> Q x) \<longleftrightarrow> (\<exists>\<^sub>\<infinity>x. P x) \<or> (\<exists>\<^sub>\<infinity>x. Q x)" by (fact frequently_disj_iff) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
201 |
lemma MOST_rev_mp: "\<forall>\<^sub>\<infinity>x. P x \<Longrightarrow> \<forall>\<^sub>\<infinity>x. P x \<longrightarrow> Q x \<Longrightarrow> \<forall>\<^sub>\<infinity>x. Q x" by (fact eventually_rev_mp) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
202 |
lemma MOST_conj_distrib: "(\<forall>\<^sub>\<infinity>x. P x \<and> Q x) \<longleftrightarrow> (\<forall>\<^sub>\<infinity>x. P x) \<and> (\<forall>\<^sub>\<infinity>x. Q x)" by (fact eventually_conj_iff) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
203 |
lemma MOST_conjI: "MOST x. P x \<Longrightarrow> MOST x. Q x \<Longrightarrow> MOST x. P x \<and> Q x" by (fact eventually_conj) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
204 |
lemma INFM_finite_Bex_distrib: "finite A \<Longrightarrow> (INFM y. \<exists>x\<in>A. P x y) \<longleftrightarrow> (\<exists>x\<in>A. INFM y. P x y)" by (fact frequently_bex_finite_distrib) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
205 |
lemma MOST_finite_Ball_distrib: "finite A \<Longrightarrow> (MOST y. \<forall>x\<in>A. P x y) \<longleftrightarrow> (\<forall>x\<in>A. MOST y. P x y)" by (fact eventually_ball_finite_distrib) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
206 |
lemma INFM_E: "INFM x. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> thesis) \<Longrightarrow> thesis" by (fact frequentlyE) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
207 |
lemma MOST_I: "(\<And>x. P x) \<Longrightarrow> MOST x. P x" by (rule eventuallyI) |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
208 |
lemmas MOST_iff_finiteNeg = MOST_iff_cofinite |
20809 | 209 |
|
210 |
||
64967 | 211 |
subsection \<open>Enumeration of an Infinite Set\<close> |
212 |
||
213 |
text \<open>The set's element type must be wellordered (e.g. the natural numbers).\<close> |
|
20809 | 214 |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
215 |
text \<open> |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
216 |
Could be generalized to |
64967 | 217 |
@{prop "enumerate' S n = (SOME t. t \<in> s \<and> finite {s\<in>S. s < t} \<and> card {s\<in>S. s < t} = n)"}. |
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
218 |
\<close> |
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
219 |
|
53239 | 220 |
primrec (in wellorder) enumerate :: "'a set \<Rightarrow> nat \<Rightarrow> 'a" |
64967 | 221 |
where |
222 |
enumerate_0: "enumerate S 0 = (LEAST n. n \<in> S)" |
|
223 |
| enumerate_Suc: "enumerate S (Suc n) = enumerate (S - {LEAST n. n \<in> S}) n" |
|
20809 | 224 |
|
53239 | 225 |
lemma enumerate_Suc': "enumerate S (Suc n) = enumerate (S - {enumerate S 0}) n" |
20809 | 226 |
by simp |
227 |
||
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
228 |
lemma enumerate_in_set: "infinite S \<Longrightarrow> enumerate S n \<in> S" |
64967 | 229 |
proof (induct n arbitrary: S) |
230 |
case 0 |
|
231 |
then show ?case |
|
232 |
by (fastforce intro: LeastI dest!: infinite_imp_nonempty) |
|
233 |
next |
|
234 |
case (Suc n) |
|
235 |
then show ?case |
|
236 |
by simp (metis DiffE infinite_remove) |
|
237 |
qed |
|
20809 | 238 |
|
239 |
declare enumerate_0 [simp del] enumerate_Suc [simp del] |
|
240 |
||
241 |
lemma enumerate_step: "infinite S \<Longrightarrow> enumerate S n < enumerate S (Suc n)" |
|
242 |
apply (induct n arbitrary: S) |
|
243 |
apply (rule order_le_neq_trans) |
|
244 |
apply (simp add: enumerate_0 Least_le enumerate_in_set) |
|
245 |
apply (simp only: enumerate_Suc') |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
246 |
apply (subgoal_tac "enumerate (S - {enumerate S 0}) 0 \<in> S - {enumerate S 0}") |
20809 | 247 |
apply (blast intro: sym) |
248 |
apply (simp add: enumerate_in_set del: Diff_iff) |
|
249 |
apply (simp add: enumerate_Suc') |
|
250 |
done |
|
251 |
||
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59506
diff
changeset
|
252 |
lemma enumerate_mono: "m < n \<Longrightarrow> infinite S \<Longrightarrow> enumerate S m < enumerate S n" |
64967 | 253 |
by (induct m n rule: less_Suc_induct) (auto intro: enumerate_step) |
20809 | 254 |
|
50134 | 255 |
lemma le_enumerate: |
256 |
assumes S: "infinite S" |
|
257 |
shows "n \<le> enumerate S n" |
|
61810 | 258 |
using S |
50134 | 259 |
proof (induct n) |
53239 | 260 |
case 0 |
261 |
then show ?case by simp |
|
262 |
next |
|
50134 | 263 |
case (Suc n) |
264 |
then have "n \<le> enumerate S n" by simp |
|
60500 | 265 |
also note enumerate_mono[of n "Suc n", OF _ \<open>infinite S\<close>] |
50134 | 266 |
finally show ?case by simp |
53239 | 267 |
qed |
50134 | 268 |
|
269 |
lemma enumerate_Suc'': |
|
270 |
fixes S :: "'a::wellorder set" |
|
53239 | 271 |
assumes "infinite S" |
272 |
shows "enumerate S (Suc n) = (LEAST s. s \<in> S \<and> enumerate S n < s)" |
|
273 |
using assms |
|
50134 | 274 |
proof (induct n arbitrary: S) |
275 |
case 0 |
|
53239 | 276 |
then have "\<forall>s \<in> S. enumerate S 0 \<le> s" |
50134 | 277 |
by (auto simp: enumerate.simps intro: Least_le) |
278 |
then show ?case |
|
279 |
unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"] |
|
53239 | 280 |
by (intro arg_cong[where f = Least] ext) auto |
50134 | 281 |
next |
282 |
case (Suc n S) |
|
283 |
show ?case |
|
60500 | 284 |
using enumerate_mono[OF zero_less_Suc \<open>infinite S\<close>, of n] \<open>infinite S\<close> |
50134 | 285 |
apply (subst (1 2) enumerate_Suc') |
286 |
apply (subst Suc) |
|
64967 | 287 |
apply (use \<open>infinite S\<close> in simp) |
53239 | 288 |
apply (intro arg_cong[where f = Least] ext) |
289 |
apply (auto simp: enumerate_Suc'[symmetric]) |
|
290 |
done |
|
50134 | 291 |
qed |
292 |
||
293 |
lemma enumerate_Ex: |
|
64967 | 294 |
fixes S :: "nat set" |
295 |
assumes S: "infinite S" |
|
296 |
and s: "s \<in> S" |
|
297 |
shows "\<exists>n. enumerate S n = s" |
|
298 |
using s |
|
50134 | 299 |
proof (induct s rule: less_induct) |
300 |
case (less s) |
|
301 |
show ?case |
|
64967 | 302 |
proof (cases "\<exists>y\<in>S. y < s") |
303 |
case True |
|
50134 | 304 |
let ?y = "Max {s'\<in>S. s' < s}" |
64967 | 305 |
from True have y: "\<And>x. ?y < x \<longleftrightarrow> (\<forall>s'\<in>S. s' < s \<longrightarrow> s' < x)" |
53239 | 306 |
by (subst Max_less_iff) auto |
307 |
then have y_in: "?y \<in> {s'\<in>S. s' < s}" |
|
308 |
by (intro Max_in) auto |
|
309 |
with less.hyps[of ?y] obtain n where "enumerate S n = ?y" |
|
310 |
by auto |
|
50134 | 311 |
with S have "enumerate S (Suc n) = s" |
312 |
by (auto simp: y less enumerate_Suc'' intro!: Least_equality) |
|
64967 | 313 |
then show ?thesis by auto |
50134 | 314 |
next |
64967 | 315 |
case False |
50134 | 316 |
then have "\<forall>t\<in>S. s \<le> t" by auto |
60500 | 317 |
with \<open>s \<in> S\<close> show ?thesis |
50134 | 318 |
by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) |
319 |
qed |
|
320 |
qed |
|
321 |
||
322 |
lemma bij_enumerate: |
|
323 |
fixes S :: "nat set" |
|
324 |
assumes S: "infinite S" |
|
325 |
shows "bij_betw (enumerate S) UNIV S" |
|
326 |
proof - |
|
327 |
have "\<And>n m. n \<noteq> m \<Longrightarrow> enumerate S n \<noteq> enumerate S m" |
|
60500 | 328 |
using enumerate_mono[OF _ \<open>infinite S\<close>] by (auto simp: neq_iff) |
50134 | 329 |
then have "inj (enumerate S)" |
330 |
by (auto simp: inj_on_def) |
|
53239 | 331 |
moreover have "\<forall>s \<in> S. \<exists>i. enumerate S i = s" |
50134 | 332 |
using enumerate_Ex[OF S] by auto |
60500 | 333 |
moreover note \<open>infinite S\<close> |
50134 | 334 |
ultimately show ?thesis |
335 |
unfolding bij_betw_def by (auto intro: enumerate_in_set) |
|
336 |
qed |
|
337 |
||
64967 | 338 |
text \<open>A pair of weird and wonderful lemmas from HOL Light.\<close> |
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
339 |
lemma finite_transitivity_chain: |
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
340 |
assumes "finite A" |
64967 | 341 |
and R: "\<And>x. \<not> R x x" "\<And>x y z. \<lbrakk>R x y; R y z\<rbrakk> \<Longrightarrow> R x z" |
342 |
and A: "\<And>x. x \<in> A \<Longrightarrow> \<exists>y. y \<in> A \<and> R x y" |
|
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
343 |
shows "A = {}" |
64967 | 344 |
using \<open>finite A\<close> A |
345 |
proof (induct A) |
|
346 |
case empty |
|
347 |
then show ?case by simp |
|
348 |
next |
|
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
349 |
case (insert a A) |
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
350 |
with R show ?case |
64967 | 351 |
by (metis empty_iff insert_iff) (* somewhat slow *) |
352 |
qed |
|
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
353 |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
354 |
corollary Union_maximal_sets: |
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
355 |
assumes "finite \<F>" |
64967 | 356 |
shows "\<Union>{T \<in> \<F>. \<forall>U\<in>\<F>. \<not> T \<subset> U} = \<Union>\<F>" |
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
357 |
(is "?lhs = ?rhs") |
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
358 |
proof |
64967 | 359 |
show "?lhs \<subseteq> ?rhs" by force |
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
360 |
show "?rhs \<subseteq> ?lhs" |
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
361 |
proof (rule Union_subsetI) |
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
362 |
fix S |
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
363 |
assume "S \<in> \<F>" |
64967 | 364 |
have "{T \<in> \<F>. S \<subseteq> T} = {}" |
365 |
if "\<not> (\<exists>y. y \<in> {T \<in> \<F>. \<forall>U\<in>\<F>. \<not> T \<subset> U} \<and> S \<subseteq> y)" |
|
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
366 |
apply (rule finite_transitivity_chain [of _ "\<lambda>T U. S \<subseteq> T \<and> T \<subset> U"]) |
64967 | 367 |
apply (use assms that in auto) |
368 |
apply (blast intro: dual_order.trans psubset_imp_subset) |
|
369 |
done |
|
370 |
with \<open>S \<in> \<F>\<close> show "\<exists>y. y \<in> {T \<in> \<F>. \<forall>U\<in>\<F>. \<not> T \<subset> U} \<and> S \<subseteq> y" |
|
371 |
by blast |
|
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
372 |
qed |
64967 | 373 |
qed |
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
61945
diff
changeset
|
374 |
|
20809 | 375 |
end |