MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  konigsberg Structured version   Visualization version   GIF version

Theorem konigsberg 27722
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 27706 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.)
Hypotheses
Ref Expression
konigsberg.v 𝑉 = (0...3)
konigsberg.e 𝐸 = ⟨“{0, 1} {0, 2} {0, 3} {1, 2} {1, 2} {2, 3} {2, 3}”⟩
konigsberg.g 𝐺 = ⟨𝑉, 𝐸
Assertion
Ref Expression
konigsberg (EulerPaths‘𝐺) = ∅

Proof of Theorem konigsberg
Dummy variable 𝑥 is distinct from all other variables.
StepHypRef 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 𝐺 = ⟨𝑉, 𝐸
41, 2, 3konigsberglem5 27721 . . 3 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)})
5 elpri 4500 . . . 4 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} → ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 ∨ (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2))
6 2pos 11594 . . . . . . 7 0 < 2
7 0re 10496 . . . . . . . 8 0 ∈ ℝ
8 2re 11565 . . . . . . . 8 2 ∈ ℝ
97, 8ltnsymi 10612 . . . . . . 7 (0 < 2 → ¬ 2 < 0)
106, 9ax-mp 5 . . . . . 6 ¬ 2 < 0
11 breq2 4972 . . . . . 6 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 → (2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ↔ 2 < 0))
1210, 11mtbiri 328 . . . . 5 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 → ¬ 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}))
138ltnri 10602 . . . . . 6 ¬ 2 < 2
14 breq2 4972 . . . . . 6 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2 → (2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ↔ 2 < 2))
1513, 14mtbiri 328 . . . . 5 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2 → ¬ 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}))
1612, 15jaoi 852 . . . 4 (((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 ∨ (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2) → ¬ 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}))
175, 16syl 17 . . 3 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} → ¬ 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}))
184, 17mt2 201 . 2 ¬ (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2}
191, 2, 3konigsbergumgr 27716 . . . . 5 𝐺 ∈ UMGraph
20 umgrupgr 26575 . . . . 5 (𝐺 ∈ UMGraph → 𝐺 ∈ UPGraph)
2119, 20ax-mp 5 . . . 4 𝐺 ∈ UPGraph
223fveq2i 6548 . . . . . 6 (Vtx‘𝐺) = (Vtx‘⟨𝑉, 𝐸⟩)
231ovexi 7056 . . . . . . 7 𝑉 ∈ V
24 s7cli 14087 . . . . . . . 8 ⟨“{0, 1} {0, 2} {0, 3} {1, 2} {1, 2} {2, 3} {2, 3}”⟩ ∈ Word V
252, 24eqeltri 2881 . . . . . . 7 𝐸 ∈ Word V
26 opvtxfv 26476 . . . . . . 7 ((𝑉 ∈ V ∧ 𝐸 ∈ Word V) → (Vtx‘⟨𝑉, 𝐸⟩) = 𝑉)
2723, 25, 26mp2an 688 . . . . . 6 (Vtx‘⟨𝑉, 𝐸⟩) = 𝑉
2822, 27eqtr2i 2822 . . . . 5 𝑉 = (Vtx‘𝐺)
2928eulerpath 27706 . . . 4 ((𝐺 ∈ UPGraph ∧ (EulerPaths‘𝐺) ≠ ∅) → (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2})
3021, 29mpan 686 . . 3 ((EulerPaths‘𝐺) ≠ ∅ → (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2})
3130necon1bi 3014 . 2 (¬ (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} → (EulerPaths‘𝐺) = ∅)
3218, 31ax-mp 5 1 (EulerPaths‘𝐺) = ∅
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wo 842   = wceq 1525  wcel 2083  wne 2986  {crab 3111  Vcvv 3440  c0 4217  {cpr 4480  cop 4484   class class class wbr 4968  cfv 6232  (class class class)co 7023  0cc0 10390  1c1 10391   < clt 10528  2c2 11546  3c3 11547  ...cfz 12746  chash 13544  Word cword 13711  ⟨“cs7 14048  cdvds 15444  Vtxcvtx 26468  UPGraphcupgr 26552  UMGraphcumgr 26553  VtxDegcvtxdg 26934  EulerPathsceupth 27662
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1781  ax-4 1795  ax-5 1892  ax-6 1951  ax-7 1996  ax-8 2085  ax-9 2093  ax-10 2114  ax-11 2128  ax-12 2143  ax-13 2346  ax-ext 2771  ax-rep 5088  ax-sep 5101  ax-nul 5108  ax-pow 5164  ax-pr 5228  ax-un 7326  ax-cnex 10446  ax-resscn 10447  ax-1cn 10448  ax-icn 10449  ax-addcl 10450  ax-addrcl 10451  ax-mulcl 10452  ax-mulrcl 10453  ax-mulcom 10454  ax-addass 10455  ax-mulass 10456  ax-distr 10457  ax-i2m1 10458  ax-1ne0 10459  ax-1rid 10460  ax-rnegex 10461  ax-rrecex 10462  ax-cnre 10463  ax-pre-lttri 10464  ax-pre-lttrn 10465  ax-pre-ltadd 10466  ax-pre-mulgt0 10467  ax-pre-sup 10468
This theorem depends on definitions:  df-bi 208  df-an 397  df-or 843  df-ifp 1056  df-3or 1081  df-3an 1082  df-tru 1528  df-ex 1766  df-nf 1770  df-sb 2045  df-mo 2578  df-eu 2614  df-clab 2778  df-cleq 2790  df-clel 2865  df-nfc 2937  df-ne 2987  df-nel 3093  df-ral 3112  df-rex 3113  df-reu 3114  df-rmo 3115  df-rab 3116  df-v 3442  df-sbc 3712  df-csb 3818  df-dif 3868  df-un 3870  df-in 3872  df-ss 3880  df-pss 3882  df-nul 4218  df-if 4388  df-pw 4461  df-sn 4479  df-pr 4481  df-tp 4483  df-op 4485  df-uni 4752  df-int 4789  df-iun 4833  df-br 4969  df-opab 5031  df-mpt 5048  df-tr 5071  df-id 5355  df-eprel 5360  df-po 5369  df-so 5370  df-fr 5409  df-we 5411  df-xp 5456  df-rel 5457  df-cnv 5458  df-co 5459  df-dm 5460  df-rn 5461  df-res 5462  df-ima 5463  df-pred 6030  df-ord 6076  df-on 6077  df-lim 6078  df-suc 6079  df-iota 6196  df-fun 6234  df-fn 6235  df-f 6236  df-f1 6237  df-fo 6238  df-f1o 6239  df-fv 6240  df-riota 6984  df-ov 7026  df-oprab 7027  df-mpo 7028  df-om 7444  df-1st 7552  df-2nd 7553  df-wrecs 7805  df-recs 7867  df-rdg 7905  df-1o 7960  df-2o 7961  df-oadd 7964  df-er 8146  df-map 8265  df-pm 8266  df-en 8365  df-dom 8366  df-sdom 8367  df-fin 8368  df-sup 8759  df-inf 8760  df-dju 9183  df-card 9221  df-pnf 10530  df-mnf 10531  df-xr 10532  df-ltxr 10533  df-le 10534  df-sub 10725  df-neg 10726  df-div 11152  df-nn 11493  df-2 11554  df-3 11555  df-4 11556  df-n0 11752  df-xnn0 11822  df-z 11836  df-uz 12098  df-rp 12244  df-xadd 12362  df-fz 12747  df-fzo 12888  df-seq 13224  df-exp 13284  df-hash 13545  df-word 13712  df-concat 13773  df-s1 13798  df-s2 14050  df-s3 14051  df-s4 14052  df-s5 14053  df-s6 14054  df-s7 14055  df-cj 14296  df-re 14297  df-im 14298  df-sqrt 14432  df-abs 14433  df-dvds 15445  df-vtx 26470  df-iedg 26471  df-edg 26520  df-uhgr 26530  df-ushgr 26531  df-upgr 26554  df-umgr 26555  df-uspgr 26622  df-vtxdg 26935  df-wlks 27068  df-trls 27160  df-eupth 27663
This theorem is referenced by: (None)
  Copyright terms: Public domain W3C validator