![]() |
Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
|
Mirrors > Home > MPE Home > Th. List > carden | Structured version Visualization version GIF version |
Description: Two sets are equinumerous
iff their cardinal numbers are equal. This
important theorem expresses the essential concept behind
"cardinality" or
"size". This theorem appears as Proposition 10.10 of [TakeutiZaring]
p. 85, Theorem 7P of [Enderton] p. 197,
and Theorem 9 of [Suppes] p. 242
(among others). The Axiom of Choice is required for its proof. Related
theorems are hasheni 14397 and the finite-set-only hashen 14396.
This theorem is also known as Hume's Principle. Gottlob Frege's two-volume Grundgesetze der Arithmetik used his Basic Law V to prove this theorem. Unfortunately Basic Law V caused Frege's system to be inconsistent because it was subject to Russell's paradox (see ru 3802). Later scholars have found that Frege primarily used Basic Law V to Hume's Principle. If Basic Law V is replaced by Hume's Principle in Frege's system, much of Frege's work is restored. Grundgesetze der Arithmetik, once Basic Law V is replaced, proves "Frege's theorem" (the Peano axioms of arithmetic can be derived in second-order logic from Hume's principle). See https://plato.stanford.edu/entries/frege-theorem 3802. We take a different approach, using first-order logic and ZFC, to prove the Peano axioms of arithmetic. The theory of cardinality can also be developed without AC by introducing "card" as a primitive notion and stating this theorem as an axiom, as is done with the axiom for cardinal numbers in [Suppes] p. 111. Finally, if we allow the Axiom of Regularity, we can avoid AC by defining the cardinal number of a set as the set of all sets equinumerous to it and having the least possible rank (see karden 9964). (Contributed by NM, 22-Oct-2003.) |
Ref | Expression |
---|---|
carden | ⊢ ((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) → ((card‘𝐴) = (card‘𝐵) ↔ 𝐴 ≈ 𝐵)) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | numth3 10539 | . . . . . 6 ⊢ (𝐴 ∈ 𝐶 → 𝐴 ∈ dom card) | |
2 | 1 | ad2antrr 725 | . . . . 5 ⊢ (((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) ∧ (card‘𝐴) = (card‘𝐵)) → 𝐴 ∈ dom card) |
3 | cardid2 10022 | . . . . 5 ⊢ (𝐴 ∈ dom card → (card‘𝐴) ≈ 𝐴) | |
4 | ensym 9063 | . . . . 5 ⊢ ((card‘𝐴) ≈ 𝐴 → 𝐴 ≈ (card‘𝐴)) | |
5 | 2, 3, 4 | 3syl 18 | . . . 4 ⊢ (((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) ∧ (card‘𝐴) = (card‘𝐵)) → 𝐴 ≈ (card‘𝐴)) |
6 | simpr 484 | . . . . 5 ⊢ (((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) ∧ (card‘𝐴) = (card‘𝐵)) → (card‘𝐴) = (card‘𝐵)) | |
7 | numth3 10539 | . . . . . . 7 ⊢ (𝐵 ∈ 𝐷 → 𝐵 ∈ dom card) | |
8 | 7 | ad2antlr 726 | . . . . . 6 ⊢ (((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) ∧ (card‘𝐴) = (card‘𝐵)) → 𝐵 ∈ dom card) |
9 | 8 | cardidd 10618 | . . . . 5 ⊢ (((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) ∧ (card‘𝐴) = (card‘𝐵)) → (card‘𝐵) ≈ 𝐵) |
10 | 6, 9 | eqbrtrd 5188 | . . . 4 ⊢ (((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) ∧ (card‘𝐴) = (card‘𝐵)) → (card‘𝐴) ≈ 𝐵) |
11 | entr 9066 | . . . 4 ⊢ ((𝐴 ≈ (card‘𝐴) ∧ (card‘𝐴) ≈ 𝐵) → 𝐴 ≈ 𝐵) | |
12 | 5, 10, 11 | syl2anc 583 | . . 3 ⊢ (((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) ∧ (card‘𝐴) = (card‘𝐵)) → 𝐴 ≈ 𝐵) |
13 | 12 | ex 412 | . 2 ⊢ ((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) → ((card‘𝐴) = (card‘𝐵) → 𝐴 ≈ 𝐵)) |
14 | carden2b 10036 | . 2 ⊢ (𝐴 ≈ 𝐵 → (card‘𝐴) = (card‘𝐵)) | |
15 | 13, 14 | impbid1 225 | 1 ⊢ ((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) → ((card‘𝐴) = (card‘𝐵) ↔ 𝐴 ≈ 𝐵)) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ↔ wb 206 ∧ wa 395 = wceq 1537 ∈ wcel 2108 class class class wbr 5166 dom cdm 5700 ‘cfv 6573 ≈ cen 9000 cardccrd 10004 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1793 ax-4 1807 ax-5 1909 ax-6 1967 ax-7 2007 ax-8 2110 ax-9 2118 ax-10 2141 ax-11 2158 ax-12 2178 ax-ext 2711 ax-rep 5303 ax-sep 5317 ax-nul 5324 ax-pow 5383 ax-pr 5447 ax-un 7770 ax-ac2 10532 |
This theorem depends on definitions: df-bi 207 df-an 396 df-or 847 df-3or 1088 df-3an 1089 df-tru 1540 df-fal 1550 df-ex 1778 df-nf 1782 df-sb 2065 df-mo 2543 df-eu 2572 df-clab 2718 df-cleq 2732 df-clel 2819 df-nfc 2895 df-ne 2947 df-ral 3068 df-rex 3077 df-rmo 3388 df-reu 3389 df-rab 3444 df-v 3490 df-sbc 3805 df-csb 3922 df-dif 3979 df-un 3981 df-in 3983 df-ss 3993 df-pss 3996 df-nul 4353 df-if 4549 df-pw 4624 df-sn 4649 df-pr 4651 df-op 4655 df-uni 4932 df-int 4971 df-iun 5017 df-br 5167 df-opab 5229 df-mpt 5250 df-tr 5284 df-id 5593 df-eprel 5599 df-po 5607 df-so 5608 df-fr 5652 df-se 5653 df-we 5654 df-xp 5706 df-rel 5707 df-cnv 5708 df-co 5709 df-dm 5710 df-rn 5711 df-res 5712 df-ima 5713 df-pred 6332 df-ord 6398 df-on 6399 df-suc 6401 df-iota 6525 df-fun 6575 df-fn 6576 df-f 6577 df-f1 6578 df-fo 6579 df-f1o 6580 df-fv 6581 df-isom 6582 df-riota 7404 df-ov 7451 df-2nd 8031 df-frecs 8322 df-wrecs 8353 df-recs 8427 df-er 8763 df-en 9004 df-card 10008 df-ac 10185 |
This theorem is referenced by: cardeq0 10621 ficard 10634 |
Copyright terms: Public domain | W3C validator |