![]() |
Metamath
Proof Explorer Theorem List (p. 94 of 429) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Color key: | ![]() (1-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | axcc4dom 9301* | Relax the constraint on axcc4 9299 to dominance instead of equinumerosity. (Contributed by Mario Carneiro, 18-Jan-2014.) |
⊢ 𝐴 ∈ V & ⊢ (𝑥 = (𝑓‘𝑛) → (𝜑 ↔ 𝜓)) ⇒ ⊢ ((𝑁 ≼ ω ∧ ∀𝑛 ∈ 𝑁 ∃𝑥 ∈ 𝐴 𝜑) → ∃𝑓(𝑓:𝑁⟶𝐴 ∧ ∀𝑛 ∈ 𝑁 𝜓)) | ||
Theorem | domtriomlem 9302* | Lemma for domtriom 9303. (Contributed by Mario Carneiro, 9-Feb-2013.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 = {𝑦 ∣ (𝑦 ⊆ 𝐴 ∧ 𝑦 ≈ 𝒫 𝑛)} & ⊢ 𝐶 = (𝑛 ∈ ω ↦ ((𝑏‘𝑛) ∖ ∪ 𝑘 ∈ 𝑛 (𝑏‘𝑘))) ⇒ ⊢ (¬ 𝐴 ∈ Fin → ω ≼ 𝐴) | ||
Theorem | domtriom 9303 | Trichotomy of equinumerosity for ω, proven using countable choice. Equivalently, all Dedekind-finite sets (as in isfin4-2 9174) are finite in the usual sense and conversely. (Contributed by Mario Carneiro, 9-Feb-2013.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (ω ≼ 𝐴 ↔ ¬ 𝐴 ≺ ω) | ||
Theorem | fin41 9304 | Under countable choice, the IV-finite sets (Dedekind-finite) coincide with I-finite (finite in the usual sense) sets. (Contributed by Mario Carneiro, 16-May-2015.) |
⊢ FinIV = Fin | ||
Theorem | dominf 9305 | A nonempty set that is a subset of its union is infinite. This version is proved from ax-cc 9295. See dominfac 9433 for a version proved from ax-ac 9319. The axiom of Regularity is used for this proof, via inf3lem6 8568, and its use is necessary: otherwise the set 𝐴 = {𝐴} or 𝐴 = {∅, 𝐴} (where the second example even has nonempty well-founded part) provides a counterexample. (Contributed by Mario Carneiro, 9-Feb-2013.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ((𝐴 ≠ ∅ ∧ 𝐴 ⊆ ∪ 𝐴) → ω ≼ 𝐴) | ||
Axiom | ax-dc 9306* | Dependent Choice. Axiom DC1 of [Schechter] p. 149. This theorem is weaker than the Axiom of Choice but is stronger than Countable Choice. It shows the existence of a sequence whose values can only be shown to exist (but cannot be constructed explicitly) and also depend on earlier values in the sequence. Dependent choice is equivalent to the statement that every (nonempty) pruned tree has a branch. This axiom is redundant in ZFC; see axdc 9381. But ZF+DC is strictly weaker than ZF+AC, so this axiom provides for theorems that do not need the full power of AC. (Contributed by Mario Carneiro, 25-Jan-2013.) |
⊢ ((∃𝑦∃𝑧 𝑦𝑥𝑧 ∧ ran 𝑥 ⊆ dom 𝑥) → ∃𝑓∀𝑛 ∈ ω (𝑓‘𝑛)𝑥(𝑓‘suc 𝑛)) | ||
Theorem | dcomex 9307 | The Axiom of Dependent Choice implies Infinity, the way we have stated it. Thus, we have Inf+AC implies DC and DC implies Inf, but AC does not imply Inf. (Contributed by Mario Carneiro, 25-Jan-2013.) |
⊢ ω ∈ V | ||
Theorem | axdc2lem 9308* | Lemma for axdc2 9309. We construct a relation 𝑅 based on 𝐹 such that 𝑥𝑅𝑦 iff 𝑦 ∈ (𝐹‘𝑥), and show that the "function" described by ax-dc 9306 can be restricted so that it is a real function (since the stated properties only show that it is the superset of a function). (Contributed by Mario Carneiro, 25-Jan-2013.) (Revised by Mario Carneiro, 26-Jun-2015.) |
⊢ 𝐴 ∈ V & ⊢ 𝑅 = {〈𝑥, 𝑦〉 ∣ (𝑥 ∈ 𝐴 ∧ 𝑦 ∈ (𝐹‘𝑥))} & ⊢ 𝐺 = (𝑥 ∈ ω ↦ (ℎ‘𝑥)) ⇒ ⊢ ((𝐴 ≠ ∅ ∧ 𝐹:𝐴⟶(𝒫 𝐴 ∖ {∅})) → ∃𝑔(𝑔:ω⟶𝐴 ∧ ∀𝑘 ∈ ω (𝑔‘suc 𝑘) ∈ (𝐹‘(𝑔‘𝑘)))) | ||
Theorem | axdc2 9309* | An apparent strengthening of ax-dc 9306 (but derived from it) which shows that there is a denumerable sequence 𝑔 for any function that maps elements of a set 𝐴 to nonempty subsets of 𝐴 such that 𝑔(𝑥 + 1) ∈ 𝐹(𝑔(𝑥)) for all 𝑥 ∈ ω. The finitistic version of this can be proven by induction, but the infinite version requires this new axiom. (Contributed by Mario Carneiro, 25-Jan-2013.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ((𝐴 ≠ ∅ ∧ 𝐹:𝐴⟶(𝒫 𝐴 ∖ {∅})) → ∃𝑔(𝑔:ω⟶𝐴 ∧ ∀𝑘 ∈ ω (𝑔‘suc 𝑘) ∈ (𝐹‘(𝑔‘𝑘)))) | ||
Theorem | axdc3lem 9310* | The class 𝑆 of finite approximations to the DC sequence is a set. (We derive here the stronger statement that 𝑆 is a subset of a specific set, namely 𝒫 (ω × 𝐴).) (Unnecessary distinct variable restrictions were removed by David Abernethy, 18-Mar-2014.) (Contributed by Mario Carneiro, 27-Jan-2013.) (Revised by Mario Carneiro, 18-Mar-2014.) |
⊢ 𝐴 ∈ V & ⊢ 𝑆 = {𝑠 ∣ ∃𝑛 ∈ ω (𝑠:suc 𝑛⟶𝐴 ∧ (𝑠‘∅) = 𝐶 ∧ ∀𝑘 ∈ 𝑛 (𝑠‘suc 𝑘) ∈ (𝐹‘(𝑠‘𝑘)))} ⇒ ⊢ 𝑆 ∈ V | ||
Theorem | axdc3lem2 9311* | Lemma for axdc3 9314. We have constructed a "candidate set" 𝑆, which consists of all finite sequences 𝑠 that satisfy our property of interest, namely 𝑠(𝑥 + 1) ∈ 𝐹(𝑠(𝑥)) on its domain, but with the added constraint that 𝑠(0) = 𝐶. These sets are possible "initial segments" of the infinite sequence satisfying these constraints, but we can leverage the standard ax-dc 9306 (with no initial condition) to select a sequence of ever-lengthening finite sequences, namely (ℎ‘𝑛):𝑚⟶𝐴 (for some integer 𝑚). We let our "choice" function select a sequence whose domain is one more than the last one, and agrees with the previous one on its domain. Thus, the application of vanilla ax-dc 9306 yields a sequence of sequences whose domains increase without bound, and whose union is a function which has all the properties we want. In this lemma, we show that given the sequence ℎ, we can construct the sequence 𝑔 that we are after. (Contributed by Mario Carneiro, 30-Jan-2013.) |
⊢ 𝐴 ∈ V & ⊢ 𝑆 = {𝑠 ∣ ∃𝑛 ∈ ω (𝑠:suc 𝑛⟶𝐴 ∧ (𝑠‘∅) = 𝐶 ∧ ∀𝑘 ∈ 𝑛 (𝑠‘suc 𝑘) ∈ (𝐹‘(𝑠‘𝑘)))} & ⊢ 𝐺 = (𝑥 ∈ 𝑆 ↦ {𝑦 ∈ 𝑆 ∣ (dom 𝑦 = suc dom 𝑥 ∧ (𝑦 ↾ dom 𝑥) = 𝑥)}) ⇒ ⊢ (∃ℎ(ℎ:ω⟶𝑆 ∧ ∀𝑘 ∈ ω (ℎ‘suc 𝑘) ∈ (𝐺‘(ℎ‘𝑘))) → ∃𝑔(𝑔:ω⟶𝐴 ∧ (𝑔‘∅) = 𝐶 ∧ ∀𝑘 ∈ ω (𝑔‘suc 𝑘) ∈ (𝐹‘(𝑔‘𝑘)))) | ||
Theorem | axdc3lem3 9312* | Simple substitution lemma for axdc3 9314. (Contributed by Mario Carneiro, 27-Jan-2013.) |
⊢ 𝐴 ∈ V & ⊢ 𝑆 = {𝑠 ∣ ∃𝑛 ∈ ω (𝑠:suc 𝑛⟶𝐴 ∧ (𝑠‘∅) = 𝐶 ∧ ∀𝑘 ∈ 𝑛 (𝑠‘suc 𝑘) ∈ (𝐹‘(𝑠‘𝑘)))} & ⊢ 𝐵 ∈ V ⇒ ⊢ (𝐵 ∈ 𝑆 ↔ ∃𝑚 ∈ ω (𝐵:suc 𝑚⟶𝐴 ∧ (𝐵‘∅) = 𝐶 ∧ ∀𝑘 ∈ 𝑚 (𝐵‘suc 𝑘) ∈ (𝐹‘(𝐵‘𝑘)))) | ||
Theorem | axdc3lem4 9313* | Lemma for axdc3 9314. We have constructed a "candidate set" 𝑆, which consists of all finite sequences 𝑠 that satisfy our property of interest, namely 𝑠(𝑥 + 1) ∈ 𝐹(𝑠(𝑥)) on its domain, but with the added constraint that 𝑠(0) = 𝐶. These sets are possible "initial segments" of the infinite sequence satisfying these constraints, but we can leverage the standard ax-dc 9306 (with no initial condition) to select a sequence of ever-lengthening finite sequences, namely (ℎ‘𝑛):𝑚⟶𝐴 (for some integer 𝑚). We let our "choice" function select a sequence whose domain is one more than the last one, and agrees with the previous one on its domain. Thus, the application of vanilla ax-dc 9306 yields a sequence of sequences whose domains increase without bound, and whose union is a function which has all the properties we want. In this lemma, we show that 𝑆 is nonempty, and that 𝐺 always maps to a nonempty subset of 𝑆, so that we can apply axdc2 9309. See axdc3lem2 9311 for the rest of the proof. (Contributed by Mario Carneiro, 27-Jan-2013.) |
⊢ 𝐴 ∈ V & ⊢ 𝑆 = {𝑠 ∣ ∃𝑛 ∈ ω (𝑠:suc 𝑛⟶𝐴 ∧ (𝑠‘∅) = 𝐶 ∧ ∀𝑘 ∈ 𝑛 (𝑠‘suc 𝑘) ∈ (𝐹‘(𝑠‘𝑘)))} & ⊢ 𝐺 = (𝑥 ∈ 𝑆 ↦ {𝑦 ∈ 𝑆 ∣ (dom 𝑦 = suc dom 𝑥 ∧ (𝑦 ↾ dom 𝑥) = 𝑥)}) ⇒ ⊢ ((𝐶 ∈ 𝐴 ∧ 𝐹:𝐴⟶(𝒫 𝐴 ∖ {∅})) → ∃𝑔(𝑔:ω⟶𝐴 ∧ (𝑔‘∅) = 𝐶 ∧ ∀𝑘 ∈ ω (𝑔‘suc 𝑘) ∈ (𝐹‘(𝑔‘𝑘)))) | ||
Theorem | axdc3 9314* | Dependent Choice. Axiom DC1 of [Schechter] p. 149, with the addition of an initial value 𝐶. This theorem is weaker than the Axiom of Choice but is stronger than Countable Choice. It shows the existence of a sequence whose values can only be shown to exist (but cannot be constructed explicitly) and also depend on earlier values in the sequence. (Contributed by Mario Carneiro, 27-Jan-2013.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ((𝐶 ∈ 𝐴 ∧ 𝐹:𝐴⟶(𝒫 𝐴 ∖ {∅})) → ∃𝑔(𝑔:ω⟶𝐴 ∧ (𝑔‘∅) = 𝐶 ∧ ∀𝑘 ∈ ω (𝑔‘suc 𝑘) ∈ (𝐹‘(𝑔‘𝑘)))) | ||
Theorem | axdc4lem 9315* | Lemma for axdc4 9316. (Contributed by Mario Carneiro, 31-Jan-2013.) (Revised by Mario Carneiro, 16-Nov-2013.) |
⊢ 𝐴 ∈ V & ⊢ 𝐺 = (𝑛 ∈ ω, 𝑥 ∈ 𝐴 ↦ ({suc 𝑛} × (𝑛𝐹𝑥))) ⇒ ⊢ ((𝐶 ∈ 𝐴 ∧ 𝐹:(ω × 𝐴)⟶(𝒫 𝐴 ∖ {∅})) → ∃𝑔(𝑔:ω⟶𝐴 ∧ (𝑔‘∅) = 𝐶 ∧ ∀𝑘 ∈ ω (𝑔‘suc 𝑘) ∈ (𝑘𝐹(𝑔‘𝑘)))) | ||
Theorem | axdc4 9316* | A more general version of axdc3 9314 that allows the function 𝐹 to vary with 𝑘. (Contributed by Mario Carneiro, 31-Jan-2013.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ((𝐶 ∈ 𝐴 ∧ 𝐹:(ω × 𝐴)⟶(𝒫 𝐴 ∖ {∅})) → ∃𝑔(𝑔:ω⟶𝐴 ∧ (𝑔‘∅) = 𝐶 ∧ ∀𝑘 ∈ ω (𝑔‘suc 𝑘) ∈ (𝑘𝐹(𝑔‘𝑘)))) | ||
Theorem | axcclem 9317* | Lemma for axcc 9318. (Contributed by Mario Carneiro, 2-Feb-2013.) (Revised by Mario Carneiro, 16-Nov-2013.) |
⊢ 𝐴 = (𝑥 ∖ {∅}) & ⊢ 𝐹 = (𝑛 ∈ ω, 𝑦 ∈ ∪ 𝐴 ↦ (𝑓‘𝑛)) & ⊢ 𝐺 = (𝑤 ∈ 𝐴 ↦ (ℎ‘suc (◡𝑓‘𝑤))) ⇒ ⊢ (𝑥 ≈ ω → ∃𝑔∀𝑧 ∈ 𝑥 (𝑧 ≠ ∅ → (𝑔‘𝑧) ∈ 𝑧)) | ||
Theorem | axcc 9318* | Although CC can be proven trivially using ac5 9337, we prove it here using DC. (New usage is discouraged.) (Contributed by Mario Carneiro, 2-Feb-2013.) |
⊢ (𝑥 ≈ ω → ∃𝑓∀𝑧 ∈ 𝑥 (𝑧 ≠ ∅ → (𝑓‘𝑧) ∈ 𝑧)) | ||
Axiom | ax-ac 9319* |
Axiom of Choice. The Axiom of Choice (AC) is usually considered an
extension of ZF set theory rather than a proper part of it. It is
sometimes considered philosophically controversial because it asserts
the existence of a set without telling us what the set is. ZF set
theory that includes AC is called ZFC.
The unpublished version given here says that given any set 𝑥, there exists a 𝑦 that is a collection of unordered pairs, one pair for each nonempty member of 𝑥. One entry in the pair is the member of 𝑥, and the other entry is some arbitrary member of that member of 𝑥. See the rewritten version ac3 9322 for a more detailed explanation. Theorem ac2 9321 shows an equivalent written compactly with restricted quantifiers. This version was specifically crafted to be short when expanded to primitives. Kurt Maes' 5-quantifier version ackm 9325 is slightly shorter when the biconditional of ax-ac 9319 is expanded into implication and negation. In axac3 9324 we allow the constant CHOICE to represent the Axiom of Choice; this simplifies the representation of theorems like gchac 9541 (the Generalized Continuum Hypothesis implies the Axiom of Choice). Standard textbook versions of AC are derived as ac8 9352, ac5 9337, and ac7 9333. The Axiom of Regularity ax-reg 8538 (among others) is used to derive our version from the standard ones; this reverse derivation is shown as theorem dfac2 8991. Equivalents to AC are the well-ordering theorem weth 9355 and Zorn's lemma zorn 9367. See ac4 9335 for comments about stronger versions of AC. In order to avoid uses of ax-reg 8538 for derivation of AC equivalents, we provide ax-ac2 9323 (due to Kurt Maes), which is equivalent to the standard AC of textbooks. The derivation of ax-ac2 9323 from ax-ac 9319 is shown by theorem axac2 9326, and the reverse derivation by axac 9327. Therefore, new proofs should normally use ax-ac2 9323 instead. (New usage is discouraged.) (Contributed by NM, 18-Jul-1996.) |
⊢ ∃𝑦∀𝑧∀𝑤((𝑧 ∈ 𝑤 ∧ 𝑤 ∈ 𝑥) → ∃𝑣∀𝑢(∃𝑡((𝑢 ∈ 𝑤 ∧ 𝑤 ∈ 𝑡) ∧ (𝑢 ∈ 𝑡 ∧ 𝑡 ∈ 𝑦)) ↔ 𝑢 = 𝑣)) | ||
Theorem | zfac 9320* | Axiom of Choice expressed with the fewest number of different variables. The penultimate step shows the logical equivalence to ax-ac 9319. (New usage is discouraged.) (Contributed by NM, 14-Aug-2003.) |
⊢ ∃𝑥∀𝑦∀𝑧((𝑦 ∈ 𝑧 ∧ 𝑧 ∈ 𝑤) → ∃𝑤∀𝑦(∃𝑤((𝑦 ∈ 𝑧 ∧ 𝑧 ∈ 𝑤) ∧ (𝑦 ∈ 𝑤 ∧ 𝑤 ∈ 𝑥)) ↔ 𝑦 = 𝑤)) | ||
Theorem | ac2 9321* | Axiom of Choice equivalent. By using restricted quantifiers, we can express the Axiom of Choice with a single explicit conjunction. (If you want to figure it out, the rewritten equivalent ac3 9322 is easier to understand.) Note: aceq0 8979 shows the logical equivalence to ax-ac 9319. (New usage is discouraged.) (Contributed by NM, 18-Jul-1996.) |
⊢ ∃𝑦∀𝑧 ∈ 𝑥 ∀𝑤 ∈ 𝑧 ∃!𝑣 ∈ 𝑧 ∃𝑢 ∈ 𝑦 (𝑧 ∈ 𝑢 ∧ 𝑣 ∈ 𝑢) | ||
Theorem | ac3 9322* |
Axiom of Choice using abbreviations. The logical equivalence to ax-ac 9319
can be established by chaining aceq0 8979 and aceq2 8980. A standard
textbook version of AC is derived from this one in dfac2a 8990, and this
version of AC is derived from the textbook version in dfac2 8991.
The following sketch will help you understand this version of the axiom. Given any set 𝑥, the axiom says that there exists a 𝑦 that is a collection of unordered pairs, one pair for each nonempty member of 𝑥. One entry in the pair is the member of 𝑥, and the other entry is some arbitrary member of that member of 𝑥. Using the Axiom of Regularity, we can show that 𝑦 is really a set of ordered pairs, very similar to the ordered pair construction opthreg 8553. The key theorem for this (used in the proof of dfac2 8991) is preleq 8552. With this modified definition of ordered pair, it can be seen that 𝑦 is actually a choice function on the members of 𝑥. For example, suppose 𝑥 = {{1, 2}, {1, 3}, {2, 3, 4}}. Let us try 𝑦 = {{{1, 2}, 1}, {{1, 3}, 1}, {{2, 3, 4}, 2}}. For the member (of 𝑥) 𝑧 = {1, 2}, the only assignment to 𝑤 and 𝑣 that satisfies the axiom is 𝑤 = 1 and 𝑣 = {{1, 2}, 1}, so there is exactly one 𝑤 as required. We verify the other two members of 𝑥 similarly. Thus, 𝑦 satisfies the axiom. Using our modified ordered pair definition, we can say that 𝑦 corresponds to the choice function {〈{1, 2}, 1〉, 〈{1, 3}, 1〉, 〈{2, 3, 4}, 2〉}. Of course other choices for 𝑦 will also satisfy the axiom, for example 𝑦 = {{{1, 2}, 2}, {{1, 3}, 1}, {{2, 3, 4}, 4}}. What AC tells us is that there exists at least one such 𝑦, but it doesn't tell us which one. (New usage is discouraged.) (Contributed by NM, 19-Jul-1996.) |
⊢ ∃𝑦∀𝑧 ∈ 𝑥 (𝑧 ≠ ∅ → ∃!𝑤 ∈ 𝑧 ∃𝑣 ∈ 𝑦 (𝑧 ∈ 𝑣 ∧ 𝑤 ∈ 𝑣)) | ||
Axiom | ax-ac2 9323* | In order to avoid uses of ax-reg 8538 for derivation of AC equivalents, we provide ax-ac2 9323, which is equivalent to the standard AC of textbooks. This appears to be the shortest known equivalent to the standard AC when expressed in terms of set theory primitives. It was found by Kurt Maes as theorem ackm 9325. We removed the leading quantifier to make it slightly shorter, since we have ax-gen 1762 available. The derivation of ax-ac2 9323 from ax-ac 9319 is shown by theorem axac2 9326, and the reverse derivation by axac 9327. Note that we use ax-reg 8538 to derive ax-ac 9319 from ax-ac2 9323, but not to derive ax-ac2 9323 from ax-ac 9319. (Contributed by NM, 19-Dec-2016.) |
⊢ ∃𝑦∀𝑧∃𝑣∀𝑢((𝑦 ∈ 𝑥 ∧ (𝑧 ∈ 𝑦 → ((𝑣 ∈ 𝑥 ∧ ¬ 𝑦 = 𝑣) ∧ 𝑧 ∈ 𝑣))) ∨ (¬ 𝑦 ∈ 𝑥 ∧ (𝑧 ∈ 𝑥 → ((𝑣 ∈ 𝑧 ∧ 𝑣 ∈ 𝑦) ∧ ((𝑢 ∈ 𝑧 ∧ 𝑢 ∈ 𝑦) → 𝑢 = 𝑣))))) | ||
Theorem | axac3 9324 | This theorem asserts that the constant CHOICE is a theorem, thus eliminating it as a hypothesis while assuming ax-ac2 9323 as an axiom. (Contributed by Mario Carneiro, 6-May-2015.) (Revised by NM, 20-Dec-2016.) (Proof modification is discouraged.) |
⊢ CHOICE | ||
Theorem | ackm 9325* |
A remarkable equivalent to the Axiom of Choice that has only five
quantifiers (when expanded to ∈, = primitives in prenex form),
discovered and proved by Kurt Maes. This establishes a new record,
reducing from 6 to 5 the largest number of quantified variables needed
by any ZFC axiom. The ZF-equivalence to AC is shown by theorem
dfackm 9026. Maes found this version of AC in April,
2004 (replacing a
longer version, also with five quantifiers, that he found in November,
2003). See Kurt Maes, "A 5-quantifier (∈,=)-expression
ZF-equivalent to the Axiom of Choice"
(http://arxiv.org/PS_cache/arxiv/pdf/0705/0705.3162v1.pdf).
The original FOM posts are: http://www.cs.nyu.edu/pipermail/fom/2003-November/007631.html http://www.cs.nyu.edu/pipermail/fom/2003-November/007641.html. (Contributed by NM, 29-Apr-2004.) (Revised by Mario Carneiro, 17-May-2015.) (Proof modification is discouraged.) |
⊢ ∀𝑥∃𝑦∀𝑧∃𝑣∀𝑢((𝑦 ∈ 𝑥 ∧ (𝑧 ∈ 𝑦 → ((𝑣 ∈ 𝑥 ∧ ¬ 𝑦 = 𝑣) ∧ 𝑧 ∈ 𝑣))) ∨ (¬ 𝑦 ∈ 𝑥 ∧ (𝑧 ∈ 𝑥 → ((𝑣 ∈ 𝑧 ∧ 𝑣 ∈ 𝑦) ∧ ((𝑢 ∈ 𝑧 ∧ 𝑢 ∈ 𝑦) → 𝑢 = 𝑣))))) | ||
Theorem | axac2 9326* | Derive ax-ac2 9323 from ax-ac 9319. (Contributed by NM, 19-Dec-2016.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ ∃𝑦∀𝑧∃𝑣∀𝑢((𝑦 ∈ 𝑥 ∧ (𝑧 ∈ 𝑦 → ((𝑣 ∈ 𝑥 ∧ ¬ 𝑦 = 𝑣) ∧ 𝑧 ∈ 𝑣))) ∨ (¬ 𝑦 ∈ 𝑥 ∧ (𝑧 ∈ 𝑥 → ((𝑣 ∈ 𝑧 ∧ 𝑣 ∈ 𝑦) ∧ ((𝑢 ∈ 𝑧 ∧ 𝑢 ∈ 𝑦) → 𝑢 = 𝑣))))) | ||
Theorem | axac 9327* | Derive ax-ac 9319 from ax-ac2 9323. Note that ax-reg 8538 is used by the proof. (Contributed by NM, 19-Dec-2016.) (Proof modification is discouraged.) |
⊢ ∃𝑦∀𝑧∀𝑤((𝑧 ∈ 𝑤 ∧ 𝑤 ∈ 𝑥) → ∃𝑣∀𝑢(∃𝑡((𝑢 ∈ 𝑤 ∧ 𝑤 ∈ 𝑡) ∧ (𝑢 ∈ 𝑡 ∧ 𝑡 ∈ 𝑦)) ↔ 𝑢 = 𝑣)) | ||
Theorem | axaci 9328 | Apply a choice equivalent. (Contributed by Mario Carneiro, 17-May-2015.) |
⊢ (CHOICE ↔ ∀𝑥𝜑) ⇒ ⊢ 𝜑 | ||
Theorem | cardeqv 9329 | All sets are well-orderable under choice. (Contributed by Mario Carneiro, 28-Apr-2015.) |
⊢ dom card = V | ||
Theorem | numth3 9330 | All sets are well-orderable under choice. (Contributed by Stefan O'Rear, 28-Feb-2015.) |
⊢ (𝐴 ∈ 𝑉 → 𝐴 ∈ dom card) | ||
Theorem | numth2 9331* | Numeration theorem: any set is equinumerous to some ordinal (using AC). Theorem 10.3 of [TakeutiZaring] p. 84. (Contributed by NM, 20-Oct-2003.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ∃𝑥 ∈ On 𝑥 ≈ 𝐴 | ||
Theorem | numth 9332* | Numeration theorem: every set can be put into one-to-one correspondence with some ordinal (using AC). Theorem 10.3 of [TakeutiZaring] p. 84. (Contributed by NM, 10-Feb-1997.) (Proof shortened by Mario Carneiro, 8-Jan-2015.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ∃𝑥 ∈ On ∃𝑓 𝑓:𝑥–1-1-onto→𝐴 | ||
Theorem | ac7 9333* | An Axiom of Choice equivalent similar to the Axiom of Choice (first form) of [Enderton] p. 49. (Contributed by NM, 29-Apr-2004.) |
⊢ ∃𝑓(𝑓 ⊆ 𝑥 ∧ 𝑓 Fn dom 𝑥) | ||
Theorem | ac7g 9334* | An Axiom of Choice equivalent similar to the Axiom of Choice (first form) of [Enderton] p. 49. (Contributed by NM, 23-Jul-2004.) |
⊢ (𝑅 ∈ 𝐴 → ∃𝑓(𝑓 ⊆ 𝑅 ∧ 𝑓 Fn dom 𝑅)) | ||
Theorem | ac4 9335* |
Equivalent of Axiom of Choice. We do not insist that 𝑓 be a
function. However, theorem ac5 9337, derived from this one, shows that
this form of the axiom does imply that at least one such set 𝑓 whose
existence we assert is in fact a function. Axiom of Choice of
[TakeutiZaring] p. 83.
Takeuti and Zaring call this "weak choice" in contrast to "strong choice" ∃𝐹∀𝑧(𝑧 ≠ ∅ → (𝐹‘𝑧) ∈ 𝑧), which asserts the existence of a universal choice function but requires second-order quantification on (proper) class variable 𝐹 and thus cannot be expressed in our first-order formalization. However, it has been shown that ZF plus strong choice is a conservative extension of ZF plus weak choice. See Ulrich Felgner, "Comparison of the axioms of local and universal choice," Fundamenta Mathematica, 71, 43-62 (1971). Weak choice can be strengthened in a different direction to choose from a collection of proper classes; see ac6s5 9351. (Contributed by NM, 21-Jul-1996.) |
⊢ ∃𝑓∀𝑧 ∈ 𝑥 (𝑧 ≠ ∅ → (𝑓‘𝑧) ∈ 𝑧) | ||
Theorem | ac4c 9336* | Equivalent of Axiom of Choice (class version). (Contributed by NM, 10-Feb-1997.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ∃𝑓∀𝑥 ∈ 𝐴 (𝑥 ≠ ∅ → (𝑓‘𝑥) ∈ 𝑥) | ||
Theorem | ac5 9337* | An Axiom of Choice equivalent: there exists a function 𝑓 (called a choice function) with domain 𝐴 that maps each nonempty member of the domain to an element of that member. Axiom AC of [BellMachover] p. 488. Note that the assertion that 𝑓 be a function is not necessary; see ac4 9335. (Contributed by NM, 29-Aug-1999.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ∃𝑓(𝑓 Fn 𝐴 ∧ ∀𝑥 ∈ 𝐴 (𝑥 ≠ ∅ → (𝑓‘𝑥) ∈ 𝑥)) | ||
Theorem | ac5b 9338* | Equivalent of Axiom of Choice. (Contributed by NM, 31-Aug-1999.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝑥 ≠ ∅ → ∃𝑓(𝑓:𝐴⟶∪ 𝐴 ∧ ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) ∈ 𝑥)) | ||
Theorem | ac6num 9339* | A version of ac6 9340 which takes the choice as a hypothesis. (Contributed by Mario Carneiro, 27-Aug-2015.) |
⊢ (𝑦 = (𝑓‘𝑥) → (𝜑 ↔ 𝜓)) ⇒ ⊢ ((𝐴 ∈ 𝑉 ∧ ∪ 𝑥 ∈ 𝐴 {𝑦 ∈ 𝐵 ∣ 𝜑} ∈ dom card ∧ ∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝜑) → ∃𝑓(𝑓:𝐴⟶𝐵 ∧ ∀𝑥 ∈ 𝐴 𝜓)) | ||
Theorem | ac6 9340* | Equivalent of Axiom of Choice. This is useful for proving that there exists, for example, a sequence mapping natural numbers to members of a larger set 𝐵, where 𝜑 depends on 𝑥 (the natural number) and 𝑦 (to specify a member of 𝐵). A stronger version of this theorem, ac6s 9344, allows 𝐵 to be a proper class. (Contributed by NM, 18-Oct-1999.) (Revised by Mario Carneiro, 27-Aug-2015.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ (𝑦 = (𝑓‘𝑥) → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝜑 → ∃𝑓(𝑓:𝐴⟶𝐵 ∧ ∀𝑥 ∈ 𝐴 𝜓)) | ||
Theorem | ac6c4 9341* | Equivalent of Axiom of Choice. 𝐵 is a collection 𝐵(𝑥) of nonempty sets. (Contributed by Mario Carneiro, 22-Mar-2013.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ≠ ∅ → ∃𝑓(𝑓 Fn 𝐴 ∧ ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) ∈ 𝐵)) | ||
Theorem | ac6c5 9342* | Equivalent of Axiom of Choice. 𝐵 is a collection 𝐵(𝑥) of nonempty sets. Remark after Theorem 10.46 of [TakeutiZaring] p. 98. (Contributed by Mario Carneiro, 22-Mar-2013.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ≠ ∅ → ∃𝑓∀𝑥 ∈ 𝐴 (𝑓‘𝑥) ∈ 𝐵) | ||
Theorem | ac9 9343* | An Axiom of Choice equivalent: the infinite Cartesian product of nonempty classes is nonempty. Axiom of Choice (second form) of [Enderton] p. 55 and its converse. (Contributed by Mario Carneiro, 22-Mar-2013.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ≠ ∅ ↔ X𝑥 ∈ 𝐴 𝐵 ≠ ∅) | ||
Theorem | ac6s 9344* | Equivalent of Axiom of Choice. Using the Boundedness Axiom bnd2 8794, we derive this strong version of ac6 9340 that doesn't require 𝐵 to be a set. (Contributed by NM, 4-Feb-2004.) |
⊢ 𝐴 ∈ V & ⊢ (𝑦 = (𝑓‘𝑥) → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝜑 → ∃𝑓(𝑓:𝐴⟶𝐵 ∧ ∀𝑥 ∈ 𝐴 𝜓)) | ||
Theorem | ac6n 9345* | Equivalent of Axiom of Choice. Contrapositive of ac6s 9344. (Contributed by NM, 10-Jun-2007.) |
⊢ 𝐴 ∈ V & ⊢ (𝑦 = (𝑓‘𝑥) → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑓(𝑓:𝐴⟶𝐵 → ∃𝑥 ∈ 𝐴 𝜓) → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐵 𝜑) | ||
Theorem | ac6s2 9346* | Generalization of the Axiom of Choice to classes. Slightly strengthened version of ac6s3 9347. (Contributed by NM, 29-Sep-2006.) |
⊢ 𝐴 ∈ V & ⊢ (𝑦 = (𝑓‘𝑥) → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 ∈ 𝐴 ∃𝑦𝜑 → ∃𝑓(𝑓 Fn 𝐴 ∧ ∀𝑥 ∈ 𝐴 𝜓)) | ||
Theorem | ac6s3 9347* | Generalization of the Axiom of Choice to classes. Theorem 10.46 of [TakeutiZaring] p. 97. (Contributed by NM, 3-Nov-2004.) |
⊢ 𝐴 ∈ V & ⊢ (𝑦 = (𝑓‘𝑥) → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 ∈ 𝐴 ∃𝑦𝜑 → ∃𝑓∀𝑥 ∈ 𝐴 𝜓) | ||
Theorem | ac6sg 9348* | ac6s 9344 with sethood as antecedent. (Contributed by FL, 3-Aug-2009.) |
⊢ (𝑦 = (𝑓‘𝑥) → (𝜑 ↔ 𝜓)) ⇒ ⊢ (𝐴 ∈ 𝑉 → (∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝜑 → ∃𝑓(𝑓:𝐴⟶𝐵 ∧ ∀𝑥 ∈ 𝐴 𝜓))) | ||
Theorem | ac6sf 9349* | Version of ac6 9340 with bound-variable hypothesis. (Contributed by NM, 2-Mar-2008.) |
⊢ Ⅎ𝑦𝜓 & ⊢ 𝐴 ∈ V & ⊢ (𝑦 = (𝑓‘𝑥) → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝜑 → ∃𝑓(𝑓:𝐴⟶𝐵 ∧ ∀𝑥 ∈ 𝐴 𝜓)) | ||
Theorem | ac6s4 9350* | Generalization of the Axiom of Choice to proper classes. 𝐵 is a collection 𝐵(𝑥) of nonempty, possible proper classes. (Contributed by NM, 29-Sep-2006.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ≠ ∅ → ∃𝑓(𝑓 Fn 𝐴 ∧ ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) ∈ 𝐵)) | ||
Theorem | ac6s5 9351* | Generalization of the Axiom of Choice to proper classes. 𝐵 is a collection 𝐵(𝑥) of nonempty, possible proper classes. Remark after Theorem 10.46 of [TakeutiZaring] p. 98. (Contributed by NM, 27-Mar-2006.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ≠ ∅ → ∃𝑓∀𝑥 ∈ 𝐴 (𝑓‘𝑥) ∈ 𝐵) | ||
Theorem | ac8 9352* | An Axiom of Choice equivalent. Given a family 𝑥 of mutually disjoint nonempty sets, there exists a set 𝑦 containing exactly one member from each set in the family. Theorem 6M(4) of [Enderton] p. 151. (Contributed by NM, 14-May-2004.) |
⊢ ((∀𝑧 ∈ 𝑥 𝑧 ≠ ∅ ∧ ∀𝑧 ∈ 𝑥 ∀𝑤 ∈ 𝑥 (𝑧 ≠ 𝑤 → (𝑧 ∩ 𝑤) = ∅)) → ∃𝑦∀𝑧 ∈ 𝑥 ∃!𝑣 𝑣 ∈ (𝑧 ∩ 𝑦)) | ||
Theorem | ac9s 9353* | An Axiom of Choice equivalent: the infinite Cartesian product of nonempty classes is nonempty. Axiom of Choice (second form) of [Enderton] p. 55 and its converse. This is a stronger version of the axiom in Enderton, with no existence requirement for the family of classes 𝐵(𝑥) (achieved via the Collection Principle cp 8792). (Contributed by NM, 29-Sep-2006.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ≠ ∅ ↔ X𝑥 ∈ 𝐴 𝐵 ≠ ∅) | ||
Theorem | numthcor 9354* | Any set is strictly dominated by some ordinal. (Contributed by NM, 22-Oct-2003.) |
⊢ (𝐴 ∈ 𝑉 → ∃𝑥 ∈ On 𝐴 ≺ 𝑥) | ||
Theorem | weth 9355* | Well-ordering theorem: any set 𝐴 can be well-ordered. This is an equivalent of the Axiom of Choice. Theorem 6 of [Suppes] p. 242. First proved by Ernst Zermelo (the "Z" in ZFC) in 1904. (Contributed by Mario Carneiro, 5-Jan-2013.) |
⊢ (𝐴 ∈ 𝑉 → ∃𝑥 𝑥 We 𝐴) | ||
Theorem | zorn2lem1 9356* | Lemma for zorn2 9366. (Contributed by NM, 3-Apr-1997.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐹 = recs((𝑓 ∈ V ↦ (℩𝑣 ∈ 𝐶 ∀𝑢 ∈ 𝐶 ¬ 𝑢𝑤𝑣))) & ⊢ 𝐶 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ ran 𝑓 𝑔𝑅𝑧} & ⊢ 𝐷 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑥)𝑔𝑅𝑧} ⇒ ⊢ ((𝑥 ∈ On ∧ (𝑤 We 𝐴 ∧ 𝐷 ≠ ∅)) → (𝐹‘𝑥) ∈ 𝐷) | ||
Theorem | zorn2lem2 9357* | Lemma for zorn2 9366. (Contributed by NM, 3-Apr-1997.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐹 = recs((𝑓 ∈ V ↦ (℩𝑣 ∈ 𝐶 ∀𝑢 ∈ 𝐶 ¬ 𝑢𝑤𝑣))) & ⊢ 𝐶 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ ran 𝑓 𝑔𝑅𝑧} & ⊢ 𝐷 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑥)𝑔𝑅𝑧} ⇒ ⊢ ((𝑥 ∈ On ∧ (𝑤 We 𝐴 ∧ 𝐷 ≠ ∅)) → (𝑦 ∈ 𝑥 → (𝐹‘𝑦)𝑅(𝐹‘𝑥))) | ||
Theorem | zorn2lem3 9358* | Lemma for zorn2 9366. (Contributed by NM, 3-Apr-1997.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐹 = recs((𝑓 ∈ V ↦ (℩𝑣 ∈ 𝐶 ∀𝑢 ∈ 𝐶 ¬ 𝑢𝑤𝑣))) & ⊢ 𝐶 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ ran 𝑓 𝑔𝑅𝑧} & ⊢ 𝐷 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑥)𝑔𝑅𝑧} ⇒ ⊢ ((𝑅 Po 𝐴 ∧ (𝑥 ∈ On ∧ (𝑤 We 𝐴 ∧ 𝐷 ≠ ∅))) → (𝑦 ∈ 𝑥 → ¬ (𝐹‘𝑥) = (𝐹‘𝑦))) | ||
Theorem | zorn2lem4 9359* | Lemma for zorn2 9366. (Contributed by NM, 3-Apr-1997.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐹 = recs((𝑓 ∈ V ↦ (℩𝑣 ∈ 𝐶 ∀𝑢 ∈ 𝐶 ¬ 𝑢𝑤𝑣))) & ⊢ 𝐶 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ ran 𝑓 𝑔𝑅𝑧} & ⊢ 𝐷 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑥)𝑔𝑅𝑧} ⇒ ⊢ ((𝑅 Po 𝐴 ∧ 𝑤 We 𝐴) → ∃𝑥 ∈ On 𝐷 = ∅) | ||
Theorem | zorn2lem5 9360* | Lemma for zorn2 9366. (Contributed by NM, 4-Apr-1997.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐹 = recs((𝑓 ∈ V ↦ (℩𝑣 ∈ 𝐶 ∀𝑢 ∈ 𝐶 ¬ 𝑢𝑤𝑣))) & ⊢ 𝐶 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ ran 𝑓 𝑔𝑅𝑧} & ⊢ 𝐷 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑥)𝑔𝑅𝑧} & ⊢ 𝐻 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑦)𝑔𝑅𝑧} ⇒ ⊢ (((𝑤 We 𝐴 ∧ 𝑥 ∈ On) ∧ ∀𝑦 ∈ 𝑥 𝐻 ≠ ∅) → (𝐹 “ 𝑥) ⊆ 𝐴) | ||
Theorem | zorn2lem6 9361* | Lemma for zorn2 9366. (Contributed by NM, 4-Apr-1997.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐹 = recs((𝑓 ∈ V ↦ (℩𝑣 ∈ 𝐶 ∀𝑢 ∈ 𝐶 ¬ 𝑢𝑤𝑣))) & ⊢ 𝐶 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ ran 𝑓 𝑔𝑅𝑧} & ⊢ 𝐷 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑥)𝑔𝑅𝑧} & ⊢ 𝐻 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑦)𝑔𝑅𝑧} ⇒ ⊢ (𝑅 Po 𝐴 → (((𝑤 We 𝐴 ∧ 𝑥 ∈ On) ∧ ∀𝑦 ∈ 𝑥 𝐻 ≠ ∅) → 𝑅 Or (𝐹 “ 𝑥))) | ||
Theorem | zorn2lem7 9362* | Lemma for zorn2 9366. (Contributed by NM, 6-Apr-1997.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐹 = recs((𝑓 ∈ V ↦ (℩𝑣 ∈ 𝐶 ∀𝑢 ∈ 𝐶 ¬ 𝑢𝑤𝑣))) & ⊢ 𝐶 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ ran 𝑓 𝑔𝑅𝑧} & ⊢ 𝐷 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑥)𝑔𝑅𝑧} & ⊢ 𝐻 = {𝑧 ∈ 𝐴 ∣ ∀𝑔 ∈ (𝐹 “ 𝑦)𝑔𝑅𝑧} ⇒ ⊢ ((𝐴 ∈ dom card ∧ 𝑅 Po 𝐴 ∧ ∀𝑠((𝑠 ⊆ 𝐴 ∧ 𝑅 Or 𝑠) → ∃𝑎 ∈ 𝐴 ∀𝑟 ∈ 𝑠 (𝑟𝑅𝑎 ∨ 𝑟 = 𝑎))) → ∃𝑎 ∈ 𝐴 ∀𝑏 ∈ 𝐴 ¬ 𝑎𝑅𝑏) | ||
Theorem | zorn2g 9363* | Zorn's Lemma of [Monk1] p. 117. This version of zorn2 9366 avoids the Axiom of Choice by assuming that 𝐴 is well-orderable. (Contributed by NM, 6-Apr-1997.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ ((𝐴 ∈ dom card ∧ 𝑅 Po 𝐴 ∧ ∀𝑤((𝑤 ⊆ 𝐴 ∧ 𝑅 Or 𝑤) → ∃𝑥 ∈ 𝐴 ∀𝑧 ∈ 𝑤 (𝑧𝑅𝑥 ∨ 𝑧 = 𝑥))) → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ¬ 𝑥𝑅𝑦) | ||
Theorem | zorng 9364* | Zorn's Lemma. If the union of every chain (with respect to inclusion) in a set belongs to the set, then the set contains a maximal element. Theorem 6M of [Enderton] p. 151. This version of zorn 9367 avoids the Axiom of Choice by assuming that 𝐴 is well-orderable. (Contributed by NM, 12-Aug-2004.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ ((𝐴 ∈ dom card ∧ ∀𝑧((𝑧 ⊆ 𝐴 ∧ [⊊] Or 𝑧) → ∪ 𝑧 ∈ 𝐴)) → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ¬ 𝑥 ⊊ 𝑦) | ||
Theorem | zornn0g 9365* | Variant of Zorn's lemma zorng 9364 in which ∅, the union of the empty chain, is not required to be an element of 𝐴. (Contributed by Jeff Madsen, 5-Jan-2011.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ ((𝐴 ∈ dom card ∧ 𝐴 ≠ ∅ ∧ ∀𝑧((𝑧 ⊆ 𝐴 ∧ 𝑧 ≠ ∅ ∧ [⊊] Or 𝑧) → ∪ 𝑧 ∈ 𝐴)) → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ¬ 𝑥 ⊊ 𝑦) | ||
Theorem | zorn2 9366* | Zorn's Lemma of [Monk1] p. 117. This theorem is equivalent to the Axiom of Choice and states that every partially ordered set 𝐴 (with an ordering relation 𝑅) in which every totally ordered subset has an upper bound, contains at least one maximal element. The main proof consists of lemmas zorn2lem1 9356 through zorn2lem7 9362; this final piece mainly changes bound variables to eliminate the hypotheses of zorn2lem7 9362. (Contributed by NM, 6-Apr-1997.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ((𝑅 Po 𝐴 ∧ ∀𝑤((𝑤 ⊆ 𝐴 ∧ 𝑅 Or 𝑤) → ∃𝑥 ∈ 𝐴 ∀𝑧 ∈ 𝑤 (𝑧𝑅𝑥 ∨ 𝑧 = 𝑥))) → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ¬ 𝑥𝑅𝑦) | ||
Theorem | zorn 9367* | Zorn's Lemma. If the union of every chain (with respect to inclusion) in a set belongs to the set, then the set contains a maximal element. This theorem is equivalent to the Axiom of Choice. Theorem 6M of [Enderton] p. 151. See zorn2 9366 for a version with general partial orderings. (Contributed by NM, 12-Aug-2004.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (∀𝑧((𝑧 ⊆ 𝐴 ∧ [⊊] Or 𝑧) → ∪ 𝑧 ∈ 𝐴) → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ¬ 𝑥 ⊊ 𝑦) | ||
Theorem | zornn0 9368* | Variant of Zorn's lemma zorn 9367 in which ∅, the union of the empty chain, is not required to be an element of 𝐴. (Contributed by Jeff Madsen, 5-Jan-2011.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ((𝐴 ≠ ∅ ∧ ∀𝑧((𝑧 ⊆ 𝐴 ∧ 𝑧 ≠ ∅ ∧ [⊊] Or 𝑧) → ∪ 𝑧 ∈ 𝐴)) → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ¬ 𝑥 ⊊ 𝑦) | ||
Theorem | ttukeylem1 9369* | Lemma for ttukey 9378. Expand out the property of being an element of a property of finite character. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ (𝜑 → 𝐹:(card‘(∪ 𝐴 ∖ 𝐵))–1-1-onto→(∪ 𝐴 ∖ 𝐵)) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) ⇒ ⊢ (𝜑 → (𝐶 ∈ 𝐴 ↔ (𝒫 𝐶 ∩ Fin) ⊆ 𝐴)) | ||
Theorem | ttukeylem2 9370* | Lemma for ttukey 9378. A property of finite character is closed under subsets. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ (𝜑 → 𝐹:(card‘(∪ 𝐴 ∖ 𝐵))–1-1-onto→(∪ 𝐴 ∖ 𝐵)) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) ⇒ ⊢ ((𝜑 ∧ (𝐶 ∈ 𝐴 ∧ 𝐷 ⊆ 𝐶)) → 𝐷 ∈ 𝐴) | ||
Theorem | ttukeylem3 9371* | Lemma for ttukey 9378. (Contributed by Mario Carneiro, 11-May-2015.) |
⊢ (𝜑 → 𝐹:(card‘(∪ 𝐴 ∖ 𝐵))–1-1-onto→(∪ 𝐴 ∖ 𝐵)) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) & ⊢ 𝐺 = recs((𝑧 ∈ V ↦ if(dom 𝑧 = ∪ dom 𝑧, if(dom 𝑧 = ∅, 𝐵, ∪ ran 𝑧), ((𝑧‘∪ dom 𝑧) ∪ if(((𝑧‘∪ dom 𝑧) ∪ {(𝐹‘∪ dom 𝑧)}) ∈ 𝐴, {(𝐹‘∪ dom 𝑧)}, ∅))))) ⇒ ⊢ ((𝜑 ∧ 𝐶 ∈ On) → (𝐺‘𝐶) = if(𝐶 = ∪ 𝐶, if(𝐶 = ∅, 𝐵, ∪ (𝐺 “ 𝐶)), ((𝐺‘∪ 𝐶) ∪ if(((𝐺‘∪ 𝐶) ∪ {(𝐹‘∪ 𝐶)}) ∈ 𝐴, {(𝐹‘∪ 𝐶)}, ∅)))) | ||
Theorem | ttukeylem4 9372* | Lemma for ttukey 9378. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ (𝜑 → 𝐹:(card‘(∪ 𝐴 ∖ 𝐵))–1-1-onto→(∪ 𝐴 ∖ 𝐵)) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) & ⊢ 𝐺 = recs((𝑧 ∈ V ↦ if(dom 𝑧 = ∪ dom 𝑧, if(dom 𝑧 = ∅, 𝐵, ∪ ran 𝑧), ((𝑧‘∪ dom 𝑧) ∪ if(((𝑧‘∪ dom 𝑧) ∪ {(𝐹‘∪ dom 𝑧)}) ∈ 𝐴, {(𝐹‘∪ dom 𝑧)}, ∅))))) ⇒ ⊢ (𝜑 → (𝐺‘∅) = 𝐵) | ||
Theorem | ttukeylem5 9373* | Lemma for ttukey 9378. The 𝐺 function forms a (transfinitely long) chain of inclusions. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ (𝜑 → 𝐹:(card‘(∪ 𝐴 ∖ 𝐵))–1-1-onto→(∪ 𝐴 ∖ 𝐵)) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) & ⊢ 𝐺 = recs((𝑧 ∈ V ↦ if(dom 𝑧 = ∪ dom 𝑧, if(dom 𝑧 = ∅, 𝐵, ∪ ran 𝑧), ((𝑧‘∪ dom 𝑧) ∪ if(((𝑧‘∪ dom 𝑧) ∪ {(𝐹‘∪ dom 𝑧)}) ∈ 𝐴, {(𝐹‘∪ dom 𝑧)}, ∅))))) ⇒ ⊢ ((𝜑 ∧ (𝐶 ∈ On ∧ 𝐷 ∈ On ∧ 𝐶 ⊆ 𝐷)) → (𝐺‘𝐶) ⊆ (𝐺‘𝐷)) | ||
Theorem | ttukeylem6 9374* | Lemma for ttukey 9378. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ (𝜑 → 𝐹:(card‘(∪ 𝐴 ∖ 𝐵))–1-1-onto→(∪ 𝐴 ∖ 𝐵)) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) & ⊢ 𝐺 = recs((𝑧 ∈ V ↦ if(dom 𝑧 = ∪ dom 𝑧, if(dom 𝑧 = ∅, 𝐵, ∪ ran 𝑧), ((𝑧‘∪ dom 𝑧) ∪ if(((𝑧‘∪ dom 𝑧) ∪ {(𝐹‘∪ dom 𝑧)}) ∈ 𝐴, {(𝐹‘∪ dom 𝑧)}, ∅))))) ⇒ ⊢ ((𝜑 ∧ 𝐶 ∈ suc (card‘(∪ 𝐴 ∖ 𝐵))) → (𝐺‘𝐶) ∈ 𝐴) | ||
Theorem | ttukeylem7 9375* | Lemma for ttukey 9378. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ (𝜑 → 𝐹:(card‘(∪ 𝐴 ∖ 𝐵))–1-1-onto→(∪ 𝐴 ∖ 𝐵)) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) & ⊢ 𝐺 = recs((𝑧 ∈ V ↦ if(dom 𝑧 = ∪ dom 𝑧, if(dom 𝑧 = ∅, 𝐵, ∪ ran 𝑧), ((𝑧‘∪ dom 𝑧) ∪ if(((𝑧‘∪ dom 𝑧) ∪ {(𝐹‘∪ dom 𝑧)}) ∈ 𝐴, {(𝐹‘∪ dom 𝑧)}, ∅))))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ 𝐴 (𝐵 ⊆ 𝑥 ∧ ∀𝑦 ∈ 𝐴 ¬ 𝑥 ⊊ 𝑦)) | ||
Theorem | ttukey2g 9376* | The Teichmüller-Tukey Lemma ttukey 9378 with a slightly stronger conclusion: we can set up the maximal element of 𝐴 so that it also contains some given 𝐵 ∈ 𝐴 as a subset. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ ((∪ 𝐴 ∈ dom card ∧ 𝐵 ∈ 𝐴 ∧ ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) → ∃𝑥 ∈ 𝐴 (𝐵 ⊆ 𝑥 ∧ ∀𝑦 ∈ 𝐴 ¬ 𝑥 ⊊ 𝑦)) | ||
Theorem | ttukeyg 9377* | The Teichmüller-Tukey Lemma ttukey 9378 stated with the "choice" as an antecedent (the hypothesis ∪ 𝐴 ∈ dom card says that ∪ 𝐴 is well-orderable). (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ ((∪ 𝐴 ∈ dom card ∧ 𝐴 ≠ ∅ ∧ ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ¬ 𝑥 ⊊ 𝑦) | ||
Theorem | ttukey 9378* | The Teichmüller-Tukey Lemma, an Axiom of Choice equivalent. If 𝐴 is a nonempty collection of finite character, then 𝐴 has a maximal element with respect to inclusion. Here "finite character" means that 𝑥 ∈ 𝐴 iff every finite subset of 𝑥 is in 𝐴. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ((𝐴 ≠ ∅ ∧ ∀𝑥(𝑥 ∈ 𝐴 ↔ (𝒫 𝑥 ∩ Fin) ⊆ 𝐴)) → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ¬ 𝑥 ⊊ 𝑦) | ||
Theorem | axdclem 9379* | Lemma for axdc 9381. (Contributed by Mario Carneiro, 25-Jan-2013.) |
⊢ 𝐹 = (rec((𝑦 ∈ V ↦ (𝑔‘{𝑧 ∣ 𝑦𝑥𝑧})), 𝑠) ↾ ω) ⇒ ⊢ ((∀𝑦 ∈ 𝒫 dom 𝑥(𝑦 ≠ ∅ → (𝑔‘𝑦) ∈ 𝑦) ∧ ran 𝑥 ⊆ dom 𝑥 ∧ ∃𝑧(𝐹‘𝐾)𝑥𝑧) → (𝐾 ∈ ω → (𝐹‘𝐾)𝑥(𝐹‘suc 𝐾))) | ||
Theorem | axdclem2 9380* | Lemma for axdc 9381. Using the full Axiom of Choice, we can construct a choice function 𝑔 on 𝒫 dom 𝑥. From this, we can build a sequence 𝐹 starting at any value 𝑠 ∈ dom 𝑥 by repeatedly applying 𝑔 to the set (𝐹‘𝑥) (where 𝑥 is the value from the previous iteration). (Contributed by Mario Carneiro, 25-Jan-2013.) |
⊢ 𝐹 = (rec((𝑦 ∈ V ↦ (𝑔‘{𝑧 ∣ 𝑦𝑥𝑧})), 𝑠) ↾ ω) ⇒ ⊢ (∃𝑧 𝑠𝑥𝑧 → (ran 𝑥 ⊆ dom 𝑥 → ∃𝑓∀𝑛 ∈ ω (𝑓‘𝑛)𝑥(𝑓‘suc 𝑛))) | ||
Theorem | axdc 9381* | This theorem derives ax-dc 9306 using ax-ac 9319 and ax-inf 8573. Thus, AC implies DC, but not vice-versa (so that ZFC is strictly stronger than ZF+DC). (New usage is discouraged.) (Contributed by Mario Carneiro, 25-Jan-2013.) |
⊢ ((∃𝑦∃𝑧 𝑦𝑥𝑧 ∧ ran 𝑥 ⊆ dom 𝑥) → ∃𝑓∀𝑛 ∈ ω (𝑓‘𝑛)𝑥(𝑓‘suc 𝑛)) | ||
Theorem | fodom 9382 | An onto function implies dominance of domain over range. Lemma 10.20 of [Kunen] p. 30. This theorem uses the Axiom of Choice ac7g 9334. AC is not needed for finite sets - see fodomfi 8280. See also fodomnum 8918. (Contributed by NM, 23-Jul-2004.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (𝐹:𝐴–onto→𝐵 → 𝐵 ≼ 𝐴) | ||
Theorem | fodomg 9383 | An onto function implies dominance of domain over range. (Contributed by NM, 23-Jul-2004.) |
⊢ (𝐴 ∈ 𝐶 → (𝐹:𝐴–onto→𝐵 → 𝐵 ≼ 𝐴)) | ||
Theorem | dmct 9384 | The domain of a countable set is countable. (Contributed by Thierry Arnoux, 29-Dec-2016.) |
⊢ (𝐴 ≼ ω → dom 𝐴 ≼ ω) | ||
Theorem | rnct 9385 | The range of a countable set is countable. (Contributed by Thierry Arnoux, 29-Dec-2016.) |
⊢ (𝐴 ≼ ω → ran 𝐴 ≼ ω) | ||
Theorem | fodomb 9386* | Equivalence of an onto mapping and dominance for a nonempty set. Proposition 10.35 of [TakeutiZaring] p. 93. (Contributed by NM, 29-Jul-2004.) |
⊢ ((𝐴 ≠ ∅ ∧ ∃𝑓 𝑓:𝐴–onto→𝐵) ↔ (∅ ≺ 𝐵 ∧ 𝐵 ≼ 𝐴)) | ||
Theorem | wdomac 9387 | When assuming AC, weak and usual dominance coincide. It is not known if this is an AC equivalent. (Contributed by Stefan O'Rear, 11-Feb-2015.) (Revised by Mario Carneiro, 5-May-2015.) |
⊢ (𝑋 ≼* 𝑌 ↔ 𝑋 ≼ 𝑌) | ||
Theorem | brdom3 9388* | Equivalence to a dominance relation. (Contributed by NM, 27-Mar-2007.) |
⊢ 𝐵 ∈ V ⇒ ⊢ (𝐴 ≼ 𝐵 ↔ ∃𝑓(∀𝑥∃*𝑦 𝑥𝑓𝑦 ∧ ∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝑦𝑓𝑥)) | ||
Theorem | brdom5 9389* | An equivalence to a dominance relation. (Contributed by NM, 29-Mar-2007.) |
⊢ 𝐵 ∈ V ⇒ ⊢ (𝐴 ≼ 𝐵 ↔ ∃𝑓(∀𝑥 ∈ 𝐵 ∃*𝑦 𝑥𝑓𝑦 ∧ ∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝑦𝑓𝑥)) | ||
Theorem | brdom4 9390* | An equivalence to a dominance relation. (Contributed by NM, 28-Mar-2007.) (Revised by NM, 16-Jun-2017.) |
⊢ 𝐵 ∈ V ⇒ ⊢ (𝐴 ≼ 𝐵 ↔ ∃𝑓(∀𝑥 ∈ 𝐵 ∃*𝑦 ∈ 𝐴 𝑥𝑓𝑦 ∧ ∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝑦𝑓𝑥)) | ||
Theorem | brdom7disj 9391* | An equivalence to a dominance relation for disjoint sets. (Contributed by NM, 29-Mar-2007.) (Revised by NM, 16-Jun-2017.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ (𝐴 ∩ 𝐵) = ∅ ⇒ ⊢ (𝐴 ≼ 𝐵 ↔ ∃𝑓(∀𝑥 ∈ 𝐵 ∃*𝑦 ∈ 𝐴 {𝑥, 𝑦} ∈ 𝑓 ∧ ∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 {𝑦, 𝑥} ∈ 𝑓)) | ||
Theorem | brdom6disj 9392* | An equivalence to a dominance relation for disjoint sets. (Contributed by NM, 5-Apr-2007.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ (𝐴 ∩ 𝐵) = ∅ ⇒ ⊢ (𝐴 ≼ 𝐵 ↔ ∃𝑓(∀𝑥 ∈ 𝐵 ∃*𝑦{𝑥, 𝑦} ∈ 𝑓 ∧ ∀𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 {𝑦, 𝑥} ∈ 𝑓)) | ||
Theorem | fin71ac 9393 | Once we allow AC, the "strongest" definition of finite set becomes equivalent to the "weakest" and the entire hierarchy collapses. (Contributed by Stefan O'Rear, 29-Oct-2014.) |
⊢ FinVII = Fin | ||
Theorem | imadomg 9394 | An image of a function under a set is dominated by the set. Proposition 10.34 of [TakeutiZaring] p. 92. (Contributed by NM, 23-Jul-2004.) |
⊢ (𝐴 ∈ 𝐵 → (Fun 𝐹 → (𝐹 “ 𝐴) ≼ 𝐴)) | ||
Theorem | fimact 9395 | The image by a function of a countable set is countable. (Contributed by Thierry Arnoux, 27-Mar-2018.) |
⊢ ((𝐴 ≼ ω ∧ Fun 𝐹) → (𝐹 “ 𝐴) ≼ ω) | ||
Theorem | fnrndomg 9396 | The range of a function is dominated by its domain. (Contributed by NM, 1-Sep-2004.) |
⊢ (𝐴 ∈ 𝐵 → (𝐹 Fn 𝐴 → ran 𝐹 ≼ 𝐴)) | ||
Theorem | fnct 9397 | If the domain of a function is countable, the function is countable. (Contributed by Thierry Arnoux, 29-Dec-2016.) |
⊢ ((𝐹 Fn 𝐴 ∧ 𝐴 ≼ ω) → 𝐹 ≼ ω) | ||
Theorem | mptct 9398* | A countable mapping set is countable. (Contributed by Thierry Arnoux, 29-Dec-2016.) |
⊢ (𝐴 ≼ ω → (𝑥 ∈ 𝐴 ↦ 𝐵) ≼ ω) | ||
Theorem | iunfo 9399* | Existence of an onto function from a disjoint union to a union. (Contributed by Mario Carneiro, 24-Jun-2013.) (Revised by Mario Carneiro, 18-Jan-2014.) |
⊢ 𝑇 = ∪ 𝑥 ∈ 𝐴 ({𝑥} × 𝐵) ⇒ ⊢ (2nd ↾ 𝑇):𝑇–onto→∪ 𝑥 ∈ 𝐴 𝐵 | ||
Theorem | iundom2g 9400* | An upper bound for the cardinality of a disjoint indexed union, with explicit choice principles. 𝐵 depends on 𝑥 and should be thought of as 𝐵(𝑥). (Contributed by Mario Carneiro, 1-Sep-2015.) |
⊢ 𝑇 = ∪ 𝑥 ∈ 𝐴 ({𝑥} × 𝐵) & ⊢ (𝜑 → ∪ 𝑥 ∈ 𝐴 (𝐶 ↑𝑚 𝐵) ∈ AC 𝐴) & ⊢ (𝜑 → ∀𝑥 ∈ 𝐴 𝐵 ≼ 𝐶) ⇒ ⊢ (𝜑 → 𝑇 ≼ (𝐴 × 𝐶)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |