| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > konigsberg | Structured version Visualization version GIF version | ||
| Description: The Königsberg Bridge problem. If 𝐺 is the Königsberg graph, i.e. a graph on four vertices 0, 1, 2, 3, with edges {0, 1}, {0, 2}, {0, 3}, {1, 2}, {1, 2}, {2, 3}, {2, 3}, then vertices 0, 1, 3 each have degree three, and 2 has degree five, so there are four vertices of odd degree and thus by eulerpath 30222 the graph cannot have an Eulerian path. It is sufficient to show that there are 3 vertices of odd degree, since a graph having an Eulerian path can only have 0 or 2 vertices of odd degree. This is Metamath 100 proof #54. (Contributed by Mario Carneiro, 11-Mar-2015.) (Revised by Mario Carneiro, 28-Feb-2016.) (Revised by AV, 9-Mar-2021.) |
| Ref | Expression |
|---|---|
| konigsberg.v | ⊢ 𝑉 = (0...3) |
| konigsberg.e | ⊢ 𝐸 = 〈“{0, 1} {0, 2} {0, 3} {1, 2} {1, 2} {2, 3} {2, 3}”〉 |
| konigsberg.g | ⊢ 𝐺 = 〈𝑉, 𝐸〉 |
| Ref | Expression |
|---|---|
| konigsberg | ⊢ (EulerPaths‘𝐺) = ∅ |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | konigsberg.v | . . . 4 ⊢ 𝑉 = (0...3) | |
| 2 | konigsberg.e | . . . 4 ⊢ 𝐸 = 〈“{0, 1} {0, 2} {0, 3} {1, 2} {1, 2} {2, 3} {2, 3}”〉 | |
| 3 | konigsberg.g | . . . 4 ⊢ 𝐺 = 〈𝑉, 𝐸〉 | |
| 4 | 1, 2, 3 | konigsberglem5 30237 | . . 3 ⊢ 2 < (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) |
| 5 | elpri 4625 | . . . 4 ⊢ ((♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} → ((♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 ∨ (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2)) | |
| 6 | 2pos 12343 | . . . . . . 7 ⊢ 0 < 2 | |
| 7 | 0re 11237 | . . . . . . . 8 ⊢ 0 ∈ ℝ | |
| 8 | 2re 12314 | . . . . . . . 8 ⊢ 2 ∈ ℝ | |
| 9 | 7, 8 | ltnsymi 11354 | . . . . . . 7 ⊢ (0 < 2 → ¬ 2 < 0) |
| 10 | 6, 9 | ax-mp 5 | . . . . . 6 ⊢ ¬ 2 < 0 |
| 11 | breq2 5123 | . . . . . 6 ⊢ ((♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 → (2 < (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ↔ 2 < 0)) | |
| 12 | 10, 11 | mtbiri 327 | . . . . 5 ⊢ ((♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 → ¬ 2 < (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)})) |
| 13 | 8 | ltnri 11344 | . . . . . 6 ⊢ ¬ 2 < 2 |
| 14 | breq2 5123 | . . . . . 6 ⊢ ((♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2 → (2 < (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ↔ 2 < 2)) | |
| 15 | 13, 14 | mtbiri 327 | . . . . 5 ⊢ ((♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2 → ¬ 2 < (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)})) |
| 16 | 12, 15 | jaoi 857 | . . . 4 ⊢ (((♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 ∨ (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2) → ¬ 2 < (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)})) |
| 17 | 5, 16 | syl 17 | . . 3 ⊢ ((♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} → ¬ 2 < (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)})) |
| 18 | 4, 17 | mt2 200 | . 2 ⊢ ¬ (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} |
| 19 | 1, 2, 3 | konigsbergumgr 30232 | . . . . 5 ⊢ 𝐺 ∈ UMGraph |
| 20 | umgrupgr 29082 | . . . . 5 ⊢ (𝐺 ∈ UMGraph → 𝐺 ∈ UPGraph) | |
| 21 | 19, 20 | ax-mp 5 | . . . 4 ⊢ 𝐺 ∈ UPGraph |
| 22 | 3 | fveq2i 6879 | . . . . . 6 ⊢ (Vtx‘𝐺) = (Vtx‘〈𝑉, 𝐸〉) |
| 23 | 1 | ovexi 7439 | . . . . . . 7 ⊢ 𝑉 ∈ V |
| 24 | s7cli 14904 | . . . . . . . 8 ⊢ 〈“{0, 1} {0, 2} {0, 3} {1, 2} {1, 2} {2, 3} {2, 3}”〉 ∈ Word V | |
| 25 | 2, 24 | eqeltri 2830 | . . . . . . 7 ⊢ 𝐸 ∈ Word V |
| 26 | opvtxfv 28983 | . . . . . . 7 ⊢ ((𝑉 ∈ V ∧ 𝐸 ∈ Word V) → (Vtx‘〈𝑉, 𝐸〉) = 𝑉) | |
| 27 | 23, 25, 26 | mp2an 692 | . . . . . 6 ⊢ (Vtx‘〈𝑉, 𝐸〉) = 𝑉 |
| 28 | 22, 27 | eqtr2i 2759 | . . . . 5 ⊢ 𝑉 = (Vtx‘𝐺) |
| 29 | 28 | eulerpath 30222 | . . . 4 ⊢ ((𝐺 ∈ UPGraph ∧ (EulerPaths‘𝐺) ≠ ∅) → (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2}) |
| 30 | 21, 29 | mpan 690 | . . 3 ⊢ ((EulerPaths‘𝐺) ≠ ∅ → (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2}) |
| 31 | 30 | necon1bi 2960 | . 2 ⊢ (¬ (♯‘{𝑥 ∈ 𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} → (EulerPaths‘𝐺) = ∅) |
| 32 | 18, 31 | ax-mp 5 | 1 ⊢ (EulerPaths‘𝐺) = ∅ |
| Colors of variables: wff setvar class |
| Syntax hints: ¬ wn 3 ∨ wo 847 = wceq 1540 ∈ wcel 2108 ≠ wne 2932 {crab 3415 Vcvv 3459 ∅c0 4308 {cpr 4603 〈cop 4607 class class class wbr 5119 ‘cfv 6531 (class class class)co 7405 0cc0 11129 1c1 11130 < clt 11269 2c2 12295 3c3 12296 ...cfz 13524 ♯chash 14348 Word cword 14531 〈“cs7 14865 ∥ cdvds 16272 Vtxcvtx 28975 UPGraphcupgr 29059 UMGraphcumgr 29060 VtxDegcvtxdg 29445 EulerPathsceupth 30178 |
| This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1795 ax-4 1809 ax-5 1910 ax-6 1967 ax-7 2007 ax-8 2110 ax-9 2118 ax-10 2141 ax-11 2157 ax-12 2177 ax-ext 2707 ax-rep 5249 ax-sep 5266 ax-nul 5276 ax-pow 5335 ax-pr 5402 ax-un 7729 ax-cnex 11185 ax-resscn 11186 ax-1cn 11187 ax-icn 11188 ax-addcl 11189 ax-addrcl 11190 ax-mulcl 11191 ax-mulrcl 11192 ax-mulcom 11193 ax-addass 11194 ax-mulass 11195 ax-distr 11196 ax-i2m1 11197 ax-1ne0 11198 ax-1rid 11199 ax-rnegex 11200 ax-rrecex 11201 ax-cnre 11202 ax-pre-lttri 11203 ax-pre-lttrn 11204 ax-pre-ltadd 11205 ax-pre-mulgt0 11206 ax-pre-sup 11207 |
| This theorem depends on definitions: df-bi 207 df-an 396 df-or 848 df-ifp 1063 df-3or 1087 df-3an 1088 df-tru 1543 df-fal 1553 df-ex 1780 df-nf 1784 df-sb 2065 df-mo 2539 df-eu 2568 df-clab 2714 df-cleq 2727 df-clel 2809 df-nfc 2885 df-ne 2933 df-nel 3037 df-ral 3052 df-rex 3061 df-rmo 3359 df-reu 3360 df-rab 3416 df-v 3461 df-sbc 3766 df-csb 3875 df-dif 3929 df-un 3931 df-in 3933 df-ss 3943 df-pss 3946 df-nul 4309 df-if 4501 df-pw 4577 df-sn 4602 df-pr 4604 df-tp 4606 df-op 4608 df-uni 4884 df-int 4923 df-iun 4969 df-br 5120 df-opab 5182 df-mpt 5202 df-tr 5230 df-id 5548 df-eprel 5553 df-po 5561 df-so 5562 df-fr 5606 df-we 5608 df-xp 5660 df-rel 5661 df-cnv 5662 df-co 5663 df-dm 5664 df-rn 5665 df-res 5666 df-ima 5667 df-pred 6290 df-ord 6355 df-on 6356 df-lim 6357 df-suc 6358 df-iota 6484 df-fun 6533 df-fn 6534 df-f 6535 df-f1 6536 df-fo 6537 df-f1o 6538 df-fv 6539 df-riota 7362 df-ov 7408 df-oprab 7409 df-mpo 7410 df-om 7862 df-1st 7988 df-2nd 7989 df-frecs 8280 df-wrecs 8311 df-recs 8385 df-rdg 8424 df-1o 8480 df-2o 8481 df-oadd 8484 df-er 8719 df-map 8842 df-pm 8843 df-en 8960 df-dom 8961 df-sdom 8962 df-fin 8963 df-sup 9454 df-inf 9455 df-dju 9915 df-card 9953 df-pnf 11271 df-mnf 11272 df-xr 11273 df-ltxr 11274 df-le 11275 df-sub 11468 df-neg 11469 df-div 11895 df-nn 12241 df-2 12303 df-3 12304 df-4 12305 df-n0 12502 df-xnn0 12575 df-z 12589 df-uz 12853 df-rp 13009 df-xadd 13129 df-fz 13525 df-fzo 13672 df-seq 14020 df-exp 14080 df-hash 14349 df-word 14532 df-concat 14589 df-s1 14614 df-s2 14867 df-s3 14868 df-s4 14869 df-s5 14870 df-s6 14871 df-s7 14872 df-cj 15118 df-re 15119 df-im 15120 df-sqrt 15254 df-abs 15255 df-dvds 16273 df-vtx 28977 df-iedg 28978 df-edg 29027 df-uhgr 29037 df-ushgr 29038 df-upgr 29061 df-umgr 29062 df-uspgr 29129 df-vtxdg 29446 df-wlks 29579 df-trls 29672 df-eupth 30179 |
| This theorem is referenced by: (None) |
| Copyright terms: Public domain | W3C validator |