Skip to content

Commit ba1a7cc

Browse files
committed
another form of finite choice
1 parent 24c641d commit ba1a7cc

File tree

3 files changed

+127
-90
lines changed

3 files changed

+127
-90
lines changed

IsarMathLib/Cardinal_ZF.thy

+121-89
Original file line numberDiff line numberDiff line change
@@ -28,7 +28,7 @@ EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *)
2828

2929
section \<open>Cardinal numbers\<close>
3030

31-
theory Cardinal_ZF imports ZF.CardinalArith func1
31+
theory Cardinal_ZF imports ZF.CardinalArith Finite_ZF func1
3232

3333
begin
3434

@@ -408,6 +408,105 @@ definition
408408
AxiomCardinalChoiceGen ("{the axiom of}_{choice holds}") where
409409
"{the axiom of} Q {choice holds} \<equiv> Card(Q) \<and> (\<forall> M N. (M \<lesssim>Q \<and> (\<forall>t\<in>M. N`t\<noteq>0)) \<longrightarrow> (\<exists>f. f:Pi(M,\<lambda>t. N`t) \<and> (\<forall>t\<in>M. f`t\<in>N`t)))"
410410

411+
412+
text\<open>The axiom of choice holds if and only if the \<open>AxiomCardinalChoice\<close>
413+
holds for every couple of a cardinal \<open>Q\<close> and a set \<open>K\<close>.\<close>
414+
415+
lemma choice_subset_imp_choice:
416+
shows "{the axiom of} Q {choice holds} \<longleftrightarrow> (\<forall> K. {the axiom of} Q {choice holds for subsets}K)"
417+
unfolding AxiomCardinalChoice_def AxiomCardinalChoiceGen_def by blast
418+
419+
text\<open>A choice axiom for greater cardinality implies one for
420+
smaller cardinality\<close>
421+
422+
lemma greater_choice_imp_smaller_choice:
423+
assumes "Q\<lesssim>Q1" "Card(Q)"
424+
shows "{the axiom of} Q1 {choice holds} \<longrightarrow> ({the axiom of} Q {choice holds})" using assms
425+
AxiomCardinalChoiceGen_def lepoll_trans by auto
426+
427+
text\<open>If we have a surjective function from a set which is injective to a set
428+
of ordinals, then we can find an injection which goes the other way.\<close>
429+
430+
lemma surj_fun_inv:
431+
assumes "f \<in> surj(A,B)" "A\<subseteq>Q" "Ord(Q)"
432+
shows "B\<lesssim>A"
433+
proof-
434+
let ?g = "{\<langle>m,\<mu> j. j\<in>A \<and> f`(j)=m\<rangle>. m\<in>B}"
435+
have "?g:B\<rightarrow>range(?g)" using lam_is_fun_range by simp
436+
then have fun: "?g:B\<rightarrow>?g``(B)" using range_image_domain by simp
437+
from assms(2,3) have OA: "\<forall>j\<in>A. Ord(j)" using lt_def Ord_in_Ord by auto
438+
{
439+
fix x
440+
assume "x\<in>?g``(B)"
441+
then have "x\<in>range(?g)" and "\<exists>y\<in>B. \<langle>y,x\<rangle>\<in>?g" by auto
442+
then obtain y where T: "x=(\<mu> j. j\<in>A\<and> f`(j)=y)" and "y\<in>B" by auto
443+
with assms(1) OA obtain z where P: "z\<in>A \<and> f`(z)=y" "Ord(z)" unfolding surj_def
444+
by auto
445+
with T have "x\<in>A \<and> f`(x)=y" using LeastI by simp
446+
hence "x\<in>A" by simp
447+
}
448+
then have "?g``(B) \<subseteq> A" by auto
449+
with fun have fun2: "?g:B\<rightarrow>A" using fun_weaken_type by auto
450+
then have "?g\<in>inj(B,A)"
451+
proof -
452+
{
453+
fix w x
454+
assume AS: "?g`w=?g`x" "w\<in>B" "x\<in>B"
455+
from assms(1) OA AS(2,3) obtain wz xz where
456+
P1: "wz\<in>A \<and> f`(wz)=w" "Ord(wz)" and P2: "xz\<in>A \<and> f`(xz)=x" "Ord(xz)"
457+
unfolding surj_def by blast
458+
from P1 have "(\<mu> j. j\<in>A\<and> f`j=w) \<in> A \<and> f`(\<mu> j. j\<in>A\<and> f`j=w)=w"
459+
by (rule LeastI)
460+
moreover from P2 have "(\<mu> j. j\<in>A\<and> f`j=x) \<in> A \<and> f`(\<mu> j. j\<in>A\<and> f`j=x)=x"
461+
by (rule LeastI)
462+
ultimately have R: "f`(\<mu> j. j\<in>A\<and> f`j=w)=w" "f`(\<mu> j. j\<in>A\<and> f`j=x)=x"
463+
by auto
464+
from AS have "(\<mu> j. j\<in>A\<and> f`(j)=w)=(\<mu> j. j\<in>A \<and> f`(j)=x)"
465+
using apply_equality fun2 by auto
466+
hence "f`(\<mu> j. j\<in>A \<and> f`(j)=w) = f`(\<mu> j. j\<in>A \<and> f`(j)=x)" by auto
467+
with R(1) have "w = f`(\<mu> j. j\<in>A\<and> f`j=x)" by auto
468+
with R(2) have "w=x" by auto
469+
}
470+
hence "\<forall>w\<in>B. \<forall>x\<in>B. ?g`(w) = ?g`(x) \<longrightarrow> w = x"
471+
by auto
472+
with fun2 show "?g\<in>inj(B,A)" unfolding inj_def by auto
473+
qed
474+
then show ?thesis unfolding lepoll_def by auto
475+
qed
476+
477+
text\<open>The difference with the previous result is that in this one
478+
\<open>A\<close> is not a subset of an ordinal, it is only injective
479+
with one.\<close>
480+
481+
theorem surj_fun_inv_2:
482+
assumes "f:surj(A,B)" "A\<lesssim>Q" "Ord(Q)"
483+
shows "B\<lesssim>A"
484+
proof-
485+
from assms(2) obtain h where h_def: "h\<in>inj(A,Q)" using lepoll_def by auto
486+
then have bij: "h\<in>bij(A,range(h))" using inj_bij_range by auto
487+
then obtain h1 where "h1\<in>bij(range(h),A)" using bij_converse_bij by auto
488+
then have "h1 \<in> surj(range(h),A)" using bij_def by auto
489+
with assms(1) have "(f O h1)\<in>surj(range(h),B)" using comp_surj by auto
490+
moreover
491+
{
492+
fix x
493+
assume p: "x\<in>range(h)"
494+
from bij have "h\<in>surj(A,range(h))" using bij_def by auto
495+
with p obtain q where "q\<in>A" and "h`(q)=x" using surj_def by auto
496+
then have "x\<in>Q" using h_def inj_def by auto
497+
}
498+
then have "range(h)\<subseteq>Q" by auto
499+
ultimately have "B\<lesssim>range(h)" using surj_fun_inv assms(3) by auto
500+
moreover have "range(h)\<approx>A" using bij eqpoll_def eqpoll_sym by blast
501+
ultimately show "B\<lesssim>A" using lepoll_eq_trans by auto
502+
qed
503+
504+
subsection\<open>Finite choice\<close>
505+
506+
text\<open>In ZF every finite collection of non-empty sets has a choice function, i.e. a function that
507+
selects one element from each set of the collection. In this section we prove various forms of
508+
that claim.\<close>
509+
411510
text\<open>The axiom of finite choice always holds.\<close>
412511

413512
theorem finite_choice:
@@ -497,97 +596,30 @@ proof -
497596
ultimately show ?thesis by (rule nat_induct)
498597
qed
499598

500-
text\<open>The axiom of choice holds if and only if the \<open>AxiomCardinalChoice\<close>
501-
holds for every couple of a cardinal \<open>Q\<close> and a set \<open>K\<close>.\<close>
502-
503-
lemma choice_subset_imp_choice:
504-
shows "{the axiom of} Q {choice holds} \<longleftrightarrow> (\<forall> K. {the axiom of} Q {choice holds for subsets}K)"
505-
unfolding AxiomCardinalChoice_def AxiomCardinalChoiceGen_def by blast
506-
507-
text\<open>A choice axiom for greater cardinality implies one for
508-
smaller cardinality\<close>
599+
text\<open>The choice functions of a collection $\mathcal{A}$ are functions $f$ defined on
600+
$\mathcal{A}$ and valued in $\bigcup \mathcal{A}$ such that $f(A)\in A$
601+
for every $A\in \mathcal{A}$.\<close>
509602

510-
lemma greater_choice_imp_smaller_choice:
511-
assumes "Q\<lesssim>Q1" "Card(Q)"
512-
shows "{the axiom of} Q1 {choice holds} \<longrightarrow> ({the axiom of} Q {choice holds})" using assms
513-
AxiomCardinalChoiceGen_def lepoll_trans by auto
514-
515-
text\<open>If we have a surjective function from a set which is injective to a set
516-
of ordinals, then we can find an injection which goes the other way.\<close>
517-
518-
lemma surj_fun_inv:
519-
assumes "f \<in> surj(A,B)" "A\<subseteq>Q" "Ord(Q)"
520-
shows "B\<lesssim>A"
521-
proof-
522-
let ?g = "{\<langle>m,\<mu> j. j\<in>A \<and> f`(j)=m\<rangle>. m\<in>B}"
523-
have "?g:B\<rightarrow>range(?g)" using lam_is_fun_range by simp
524-
then have fun: "?g:B\<rightarrow>?g``(B)" using range_image_domain by simp
525-
from assms(2,3) have OA: "\<forall>j\<in>A. Ord(j)" using lt_def Ord_in_Ord by auto
526-
{
527-
fix x
528-
assume "x\<in>?g``(B)"
529-
then have "x\<in>range(?g)" and "\<exists>y\<in>B. \<langle>y,x\<rangle>\<in>?g" by auto
530-
then obtain y where T: "x=(\<mu> j. j\<in>A\<and> f`(j)=y)" and "y\<in>B" by auto
531-
with assms(1) OA obtain z where P: "z\<in>A \<and> f`(z)=y" "Ord(z)" unfolding surj_def
532-
by auto
533-
with T have "x\<in>A \<and> f`(x)=y" using LeastI by simp
534-
hence "x\<in>A" by simp
535-
}
536-
then have "?g``(B) \<subseteq> A" by auto
537-
with fun have fun2: "?g:B\<rightarrow>A" using fun_weaken_type by auto
538-
then have "?g\<in>inj(B,A)"
539-
proof -
540-
{
541-
fix w x
542-
assume AS: "?g`w=?g`x" "w\<in>B" "x\<in>B"
543-
from assms(1) OA AS(2,3) obtain wz xz where
544-
P1: "wz\<in>A \<and> f`(wz)=w" "Ord(wz)" and P2: "xz\<in>A \<and> f`(xz)=x" "Ord(xz)"
545-
unfolding surj_def by blast
546-
from P1 have "(\<mu> j. j\<in>A\<and> f`j=w) \<in> A \<and> f`(\<mu> j. j\<in>A\<and> f`j=w)=w"
547-
by (rule LeastI)
548-
moreover from P2 have "(\<mu> j. j\<in>A\<and> f`j=x) \<in> A \<and> f`(\<mu> j. j\<in>A\<and> f`j=x)=x"
549-
by (rule LeastI)
550-
ultimately have R: "f`(\<mu> j. j\<in>A\<and> f`j=w)=w" "f`(\<mu> j. j\<in>A\<and> f`j=x)=x"
551-
by auto
552-
from AS have "(\<mu> j. j\<in>A\<and> f`(j)=w)=(\<mu> j. j\<in>A \<and> f`(j)=x)"
553-
using apply_equality fun2 by auto
554-
hence "f`(\<mu> j. j\<in>A \<and> f`(j)=w) = f`(\<mu> j. j\<in>A \<and> f`(j)=x)" by auto
555-
with R(1) have "w = f`(\<mu> j. j\<in>A\<and> f`j=x)" by auto
556-
with R(2) have "w=x" by auto
557-
}
558-
hence "\<forall>w\<in>B. \<forall>x\<in>B. ?g`(w) = ?g`(x) \<longrightarrow> w = x"
559-
by auto
560-
with fun2 show "?g\<in>inj(B,A)" unfolding inj_def by auto
561-
qed
562-
then show ?thesis unfolding lepoll_def by auto
563-
qed
603+
definition
604+
"ChoiceFunctions(\<A>) \<equiv> {f\<in>\<A>\<rightarrow>\<Union>\<A>. \<forall>A\<in>\<A>. f`(A)\<in>A}"
564605

565-
text\<open>The difference with the previous result is that in this one
566-
\<open>A\<close> is not a subset of an ordinal, it is only injective
567-
with one.\<close>
606+
text\<open>For finite collections of non-empty sets the set of choice functions is non-empty.\<close>
568607

569-
theorem surj_fun_inv_2:
570-
assumes "f:surj(A,B)" "A\<lesssim>Q" "Ord(Q)"
571-
shows "B\<lesssim>A"
572-
proof-
573-
from assms(2) obtain h where h_def: "h\<in>inj(A,Q)" using lepoll_def by auto
574-
then have bij: "h\<in>bij(A,range(h))" using inj_bij_range by auto
575-
then obtain h1 where "h1\<in>bij(range(h),A)" using bij_converse_bij by auto
576-
then have "h1 \<in> surj(range(h),A)" using bij_def by auto
577-
with assms(1) have "(f O h1)\<in>surj(range(h),B)" using comp_surj by auto
578-
moreover
579-
{
580-
fix x
581-
assume p: "x\<in>range(h)"
582-
from bij have "h\<in>surj(A,range(h))" using bij_def by auto
583-
with p obtain q where "q\<in>A" and "h`(q)=x" using surj_def by auto
584-
then have "x\<in>Q" using h_def inj_def by auto
585-
}
586-
then have "range(h)\<subseteq>Q" by auto
587-
ultimately have "B\<lesssim>range(h)" using surj_fun_inv assms(3) by auto
588-
moreover have "range(h)\<approx>A" using bij eqpoll_def eqpoll_sym by blast
589-
ultimately show "B\<lesssim>A" using lepoll_eq_trans by auto
608+
theorem finite_choice1: assumes "Finite(\<A>)" and "\<forall>A\<in>\<A>. A\<noteq>\<emptyset>"
609+
shows "ChoiceFunctions(\<A>) \<noteq> \<emptyset>"
610+
proof -
611+
let ?N = "id(\<A>)"
612+
let ?\<N> = "(\<lambda>t. ?N`(t))"
613+
from assms(1) obtain n where "n\<in>nat" and "\<A>\<approx>n"
614+
unfolding Finite_def by auto
615+
from assms(2) \<open>\<A>\<approx>n\<close> have "\<A>\<lesssim>n" and "\<forall>A\<in>\<A>. ?N`(A)\<noteq>\<emptyset>"
616+
using eqpoll_imp_lepoll by simp_all
617+
with \<open>n\<in>nat\<close> obtain f where "f\<in>Pi(\<A>,?\<N>)"
618+
using finite_choice unfolding AxiomCardinalChoiceGen_def by blast
619+
have "Pi(\<A>,?\<N>) = {f\<in>\<A>\<rightarrow>(\<Union>A\<in>\<A>. ?\<N>(A)). \<forall>A\<in>\<A>. f`(A)\<in>?\<N>(A)}"
620+
by (rule pi_fun_space)
621+
with \<open>f\<in>Pi(\<A>,?\<N>)\<close> show ?thesis
622+
unfolding ChoiceFunctions_def by auto
590623
qed
591624

592-
593625
end

IsarMathLib/Finite_ZF.thy

+5
Original file line numberDiff line numberDiff line change
@@ -64,6 +64,11 @@ lemma card_fin_is_nat: assumes "A \<in> FinPow(X)"
6464
using assms FinPow_def Finite_def cardinal_cong nat_into_Card
6565
Card_cardinal_eq by auto
6666

67+
text\<open>The cardinality of a finite set is a natural number.\<close>
68+
69+
lemma card_fin_is_nat1: assumes "Finite(A)" shows "|A| \<in> nat"
70+
using assms card_fin_is_nat(1) unfolding FinPow_def by auto
71+
6772
text\<open>A reformulation of \<open>card_fin_is_nat\<close>: for a finit
6873
set $A$ there is a bijection between $|A|$ and $A$.\<close>
6974

IsarMathLib/func1.thy

+1-1
Original file line numberDiff line numberDiff line change
@@ -955,7 +955,7 @@ qed
955955
subsection\<open>Dependent function space\<close>
956956

957957
text\<open>The standard Isabelle/ZF \<open>ZF_Base\<close> theory defines a general notion of
958-
a dependent function space $\text{Pi}(X,N)$, where $X$ wher $X$ is any set and
958+
a dependent function space \<open>Pi(X,N)\<close>, where $X$ is any set and
959959
$N$ is a collection of sets indexed by $X$. We rarely use that notion in IsarMathLib.
960960
The facts shown in this section provide information on how to interpret the
961961
dependent function space notion in terms of a regular space of functions defined on $X$

0 commit comments

Comments
 (0)