src/ZF/AC.thy
author Fabian Huch <huch@in.tum.de>
Tue, 06 Aug 2024 15:40:51 +0200
changeset 80647 3519f026e7d6
parent 76215 a642599ffdea
permissions -rw-r--r--
tuned and clarified;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 484
diff changeset
     1
(*  Title:      ZF/AC.thy
2b8c2a7547ab expanded tabs
clasohm
parents: 484
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     3
    Copyright   1994  University of Cambridge
13328
703de709a64b better document preparation
paulson
parents: 13269
diff changeset
     4
*)
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     5
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
     6
section\<open>The Axiom of Choice\<close>
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     7
65449
c82e63b11b8b clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 61980
diff changeset
     8
theory AC imports ZF begin
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
     9
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    10
text\<open>This definition comes from Halmos (1960), page 59.\<close>
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    11
axiomatization where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 65449
diff changeset
    12
  AC: "\<lbrakk>a \<in> A;  \<And>x. x \<in> A \<Longrightarrow> (\<exists>y. y \<in> B(x))\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> Pi(A,B)"
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    13
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    14
(*The same as AC, but no premise @{term"a \<in> A"}*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 65449
diff changeset
    15
lemma AC_Pi: "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> (\<exists>y. y \<in> B(x))\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> Pi(A,B)"
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    16
apply (case_tac "A=0")
13149
773657d466cb better simplification of trivial existential equalities
paulson
parents: 13134
diff changeset
    17
apply (simp add: Pi_empty1)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    18
(*The non-trivial case*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    19
apply (blast intro: AC)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    20
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    21
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    22
(*Using dtac, this has the advantage of DELETING the universal quantifier*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 65449
diff changeset
    23
lemma AC_ball_Pi: "\<forall>x \<in> A. \<exists>y. y \<in> B(x) \<Longrightarrow> \<exists>y. y \<in> Pi(A,B)"
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    24
apply (rule AC_Pi)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13149
diff changeset
    25
apply (erule bspec, assumption)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    26
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    27
61980
6b780867d426 clarified syntax;
wenzelm
parents: 60770
diff changeset
    28
lemma AC_Pi_Pow: "\<exists>f. f \<in> (\<Prod>X \<in> Pow(C)-{0}. X)"
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
    29
apply (rule_tac B1 = "\<lambda>x. x" in AC_Pi [THEN exE])
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13149
diff changeset
    30
apply (erule_tac [2] exI, blast)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    31
done
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 2469
diff changeset
    32
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    33
lemma AC_func:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 65449
diff changeset
    34
     "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> (\<exists>y. y \<in> x)\<rbrakk> \<Longrightarrow> \<exists>f \<in> A->\<Union>(A). \<forall>x \<in> A. f`x \<in> x"
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
    35
apply (rule_tac B1 = "\<lambda>x. x" in AC_Pi [THEN exE])
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    36
prefer 2 apply (blast dest: apply_type intro: Pi_type, blast)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    37
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    38
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 65449
diff changeset
    39
lemma non_empty_family: "\<lbrakk>0 \<notin> A;  x \<in> A\<rbrakk> \<Longrightarrow> \<exists>y. y \<in> x"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13149
diff changeset
    40
by (subgoal_tac "x \<noteq> 0", blast+)
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 2469
diff changeset
    41
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 65449
diff changeset
    42
lemma AC_func0: "0 \<notin> A \<Longrightarrow> \<exists>f \<in> A->\<Union>(A). \<forall>x \<in> A. f`x \<in> x"
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    43
apply (rule AC_func)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    44
apply (simp_all add: non_empty_family)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    45
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    46
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    47
lemma AC_func_Pow: "\<exists>f \<in> (Pow(C)-{0}) -> C. \<forall>x \<in> Pow(C)-{0}. f`x \<in> x"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    48
apply (rule AC_func0 [THEN bexE])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    49
apply (rule_tac [2] bexI)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13149
diff changeset
    50
prefer 2 apply assumption
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13149
diff changeset
    51
apply (erule_tac [2] fun_weaken_type, blast+)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    52
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    53
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 65449
diff changeset
    54
lemma AC_Pi0: "0 \<notin> A \<Longrightarrow> \<exists>f. f \<in> (\<Prod>x \<in> A. x)"
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    55
apply (rule AC_Pi)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    56
apply (simp_all add: non_empty_family)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    57
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    58
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    59
end