Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
Mirrors > Home > MPE Home > Th. List > phplem3 | Structured version Visualization version GIF version |
Description: Lemma for Pigeonhole Principle. A natural number is equinumerous to its successor minus any element of the successor. (Contributed by NM, 26-May-1998.) |
Ref | Expression |
---|---|
phplem2.1 | ⊢ 𝐴 ∈ V |
phplem2.2 | ⊢ 𝐵 ∈ V |
Ref | Expression |
---|---|
phplem3 | ⊢ ((𝐴 ∈ ω ∧ 𝐵 ∈ suc 𝐴) → 𝐴 ≈ (suc 𝐴 ∖ {𝐵})) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | elsuci 6279 | . 2 ⊢ (𝐵 ∈ suc 𝐴 → (𝐵 ∈ 𝐴 ∨ 𝐵 = 𝐴)) | |
2 | phplem2.1 | . . . 4 ⊢ 𝐴 ∈ V | |
3 | phplem2.2 | . . . 4 ⊢ 𝐵 ∈ V | |
4 | 2, 3 | phplem2 8826 | . . 3 ⊢ ((𝐴 ∈ ω ∧ 𝐵 ∈ 𝐴) → 𝐴 ≈ (suc 𝐴 ∖ {𝐵})) |
5 | 2 | enref 8661 | . . . 4 ⊢ 𝐴 ≈ 𝐴 |
6 | nnord 7652 | . . . . . 6 ⊢ (𝐴 ∈ ω → Ord 𝐴) | |
7 | orddif 6306 | . . . . . 6 ⊢ (Ord 𝐴 → 𝐴 = (suc 𝐴 ∖ {𝐴})) | |
8 | 6, 7 | syl 17 | . . . . 5 ⊢ (𝐴 ∈ ω → 𝐴 = (suc 𝐴 ∖ {𝐴})) |
9 | sneq 4551 | . . . . . . 7 ⊢ (𝐴 = 𝐵 → {𝐴} = {𝐵}) | |
10 | 9 | difeq2d 4037 | . . . . . 6 ⊢ (𝐴 = 𝐵 → (suc 𝐴 ∖ {𝐴}) = (suc 𝐴 ∖ {𝐵})) |
11 | 10 | eqcoms 2745 | . . . . 5 ⊢ (𝐵 = 𝐴 → (suc 𝐴 ∖ {𝐴}) = (suc 𝐴 ∖ {𝐵})) |
12 | 8, 11 | sylan9eq 2798 | . . . 4 ⊢ ((𝐴 ∈ ω ∧ 𝐵 = 𝐴) → 𝐴 = (suc 𝐴 ∖ {𝐵})) |
13 | 5, 12 | breqtrid 5090 | . . 3 ⊢ ((𝐴 ∈ ω ∧ 𝐵 = 𝐴) → 𝐴 ≈ (suc 𝐴 ∖ {𝐵})) |
14 | 4, 13 | jaodan 958 | . 2 ⊢ ((𝐴 ∈ ω ∧ (𝐵 ∈ 𝐴 ∨ 𝐵 = 𝐴)) → 𝐴 ≈ (suc 𝐴 ∖ {𝐵})) |
15 | 1, 14 | sylan2 596 | 1 ⊢ ((𝐴 ∈ ω ∧ 𝐵 ∈ suc 𝐴) → 𝐴 ≈ (suc 𝐴 ∖ {𝐵})) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ∧ wa 399 ∨ wo 847 = wceq 1543 ∈ wcel 2110 Vcvv 3408 ∖ cdif 3863 {csn 4541 class class class wbr 5053 Ord word 6212 suc csuc 6215 ωcom 7644 ≈ cen 8623 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1803 ax-4 1817 ax-5 1918 ax-6 1976 ax-7 2016 ax-8 2112 ax-9 2120 ax-10 2141 ax-11 2158 ax-12 2175 ax-ext 2708 ax-sep 5192 ax-nul 5199 ax-pow 5258 ax-pr 5322 ax-un 7523 |
This theorem depends on definitions: df-bi 210 df-an 400 df-or 848 df-3an 1091 df-tru 1546 df-fal 1556 df-ex 1788 df-nf 1792 df-sb 2071 df-mo 2539 df-eu 2568 df-clab 2715 df-cleq 2729 df-clel 2816 df-nfc 2886 df-ne 2941 df-ral 3066 df-rex 3067 df-rab 3070 df-v 3410 df-dif 3869 df-un 3871 df-in 3873 df-ss 3883 df-nul 4238 df-if 4440 df-pw 4515 df-sn 4542 df-pr 4544 df-op 4548 df-uni 4820 df-br 5054 df-opab 5116 df-tr 5162 df-id 5455 df-eprel 5460 df-po 5468 df-so 5469 df-fr 5509 df-we 5511 df-xp 5557 df-rel 5558 df-cnv 5559 df-co 5560 df-dm 5561 df-rn 5562 df-res 5563 df-ima 5564 df-ord 6216 df-on 6217 df-suc 6219 df-fun 6382 df-fn 6383 df-f 6384 df-f1 6385 df-fo 6386 df-f1o 6387 df-om 7645 df-en 8627 |
This theorem is referenced by: phplem4 8828 php 8830 |
Copyright terms: Public domain | W3C validator |