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

Theorem konigsberg 28030
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 28014 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 28029 . . 3 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)})
5 elpri 4582 . . . 4 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} → ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 ∨ (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2))
6 2pos 11734 . . . . . . 7 0 < 2
7 0re 10637 . . . . . . . 8 0 ∈ ℝ
8 2re 11705 . . . . . . . 8 2 ∈ ℝ
97, 8ltnsymi 10753 . . . . . . 7 (0 < 2 → ¬ 2 < 0)
106, 9ax-mp 5 . . . . . 6 ¬ 2 < 0
11 breq2 5062 . . . . . 6 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 → (2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ↔ 2 < 0))
1210, 11mtbiri 329 . . . . 5 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 → ¬ 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}))
138ltnri 10743 . . . . . 6 ¬ 2 < 2
14 breq2 5062 . . . . . 6 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2 → (2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ↔ 2 < 2))
1513, 14mtbiri 329 . . . . 5 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2 → ¬ 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}))
1612, 15jaoi 853 . . . 4 (((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 0 ∨ (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) = 2) → ¬ 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}))
175, 16syl 17 . . 3 ((♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} → ¬ 2 < (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}))
184, 17mt2 202 . 2 ¬ (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2}
191, 2, 3konigsbergumgr 28024 . . . . 5 𝐺 ∈ UMGraph
20 umgrupgr 26882 . . . . 5 (𝐺 ∈ UMGraph → 𝐺 ∈ UPGraph)
2119, 20ax-mp 5 . . . 4 𝐺 ∈ UPGraph
223fveq2i 6667 . . . . . 6 (Vtx‘𝐺) = (Vtx‘⟨𝑉, 𝐸⟩)
231ovexi 7184 . . . . . . 7 𝑉 ∈ V
24 s7cli 14241 . . . . . . . 8 ⟨“{0, 1} {0, 2} {0, 3} {1, 2} {1, 2} {2, 3} {2, 3}”⟩ ∈ Word V
252, 24eqeltri 2909 . . . . . . 7 𝐸 ∈ Word V
26 opvtxfv 26783 . . . . . . 7 ((𝑉 ∈ V ∧ 𝐸 ∈ Word V) → (Vtx‘⟨𝑉, 𝐸⟩) = 𝑉)
2723, 25, 26mp2an 690 . . . . . 6 (Vtx‘⟨𝑉, 𝐸⟩) = 𝑉
2822, 27eqtr2i 2845 . . . . 5 𝑉 = (Vtx‘𝐺)
2928eulerpath 28014 . . . 4 ((𝐺 ∈ UPGraph ∧ (EulerPaths‘𝐺) ≠ ∅) → (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2})
3021, 29mpan 688 . . 3 ((EulerPaths‘𝐺) ≠ ∅ → (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2})
3130necon1bi 3044 . 2 (¬ (♯‘{𝑥𝑉 ∣ ¬ 2 ∥ ((VtxDeg‘𝐺)‘𝑥)}) ∈ {0, 2} → (EulerPaths‘𝐺) = ∅)
3218, 31ax-mp 5 1 (EulerPaths‘𝐺) = ∅
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wo 843   = wceq 1533  wcel 2110  wne 3016  {crab 3142  Vcvv 3494  c0 4290  {cpr 4562  cop 4566   class class class wbr 5058  cfv 6349  (class class class)co 7150  0cc0 10531  1c1 10532   < clt 10669  2c2 11686  3c3 11687  ...cfz 12886  chash 13684  Word cword 13855  ⟨“cs7 14202  cdvds 15601  Vtxcvtx 26775  UPGraphcupgr 26859  UMGraphcumgr 26860  VtxDegcvtxdg 27241  EulerPathsceupth 27970
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1792  ax-4 1806  ax-5 1907  ax-6 1966  ax-7 2011  ax-8 2112  ax-9 2120  ax-10 2141  ax-11 2157  ax-12 2173  ax-ext 2793  ax-rep 5182  ax-sep 5195  ax-nul 5202  ax-pow 5258  ax-pr 5321  ax-un 7455  ax-cnex 10587  ax-resscn 10588  ax-1cn 10589  ax-icn 10590  ax-addcl 10591  ax-addrcl 10592  ax-mulcl 10593  ax-mulrcl 10594  ax-mulcom 10595  ax-addass 10596  ax-mulass 10597  ax-distr 10598  ax-i2m1 10599  ax-1ne0 10600  ax-1rid 10601  ax-rnegex 10602  ax-rrecex 10603  ax-cnre 10604  ax-pre-lttri 10605  ax-pre-lttrn 10606  ax-pre-ltadd 10607  ax-pre-mulgt0 10608  ax-pre-sup 10609
This theorem depends on definitions:  df-bi 209  df-an 399  df-or 844  df-ifp 1058  df-3or 1084  df-3an 1085  df-tru 1536  df-ex 1777  df-nf 1781  df-sb 2066  df-mo 2618  df-eu 2650  df-clab 2800  df-cleq 2814  df-clel 2893  df-nfc 2963  df-ne 3017  df-nel 3124  df-ral 3143  df-rex 3144  df-reu 3145  df-rmo 3146  df-rab 3147  df-v 3496  df-sbc 3772  df-csb 3883  df-dif 3938  df-un 3940  df-in 3942  df-ss 3951  df-pss 3953  df-nul 4291  df-if 4467  df-pw 4540  df-sn 4561  df-pr 4563  df-tp 4565  df-op 4567  df-uni 4832  df-int 4869  df-iun 4913  df-br 5059  df-opab 5121  df-mpt 5139  df-tr 5165  df-id 5454  df-eprel 5459  df-po 5468  df-so 5469  df-fr 5508  df-we 5510  df-xp 5555  df-rel 5556  df-cnv 5557  df-co 5558  df-dm 5559  df-rn 5560  df-res 5561  df-ima 5562  df-pred 6142  df-ord 6188  df-on 6189  df-lim 6190  df-suc 6191  df-iota 6308  df-fun 6351  df-fn 6352  df-f 6353  df-f1 6354  df-fo 6355  df-f1o 6356  df-fv 6357  df-riota 7108  df-ov 7153  df-oprab 7154  df-mpo 7155  df-om 7575  df-1st 7683  df-2nd 7684  df-wrecs 7941  df-recs 8002  df-rdg 8040  df-1o 8096  df-2o 8097  df-oadd 8100  df-er 8283  df-map 8402  df-pm 8403  df-en 8504  df-dom 8505  df-sdom 8506  df-fin 8507  df-sup 8900  df-inf 8901  df-dju 9324  df-card 9362  df-pnf 10671  df-mnf 10672  df-xr 10673  df-ltxr 10674  df-le 10675  df-sub 10866  df-neg 10867  df-div 11292  df-nn 11633  df-2 11694  df-3 11695  df-4 11696  df-n0 11892  df-xnn0 11962  df-z 11976  df-uz 12238  df-rp 12384  df-xadd 12502  df-fz 12887  df-fzo 13028  df-seq 13364  df-exp 13424  df-hash 13685  df-word 13856  df-concat 13917  df-s1 13944  df-s2 14204  df-s3 14205  df-s4 14206  df-s5 14207  df-s6 14208  df-s7 14209  df-cj 14452  df-re 14453  df-im 14454  df-sqrt 14588  df-abs 14589  df-dvds 15602  df-vtx 26777  df-iedg 26778  df-edg 26827  df-uhgr 26837  df-ushgr 26838  df-upgr 26861  df-umgr 26862  df-uspgr 26929  df-vtxdg 27242  df-wlks 27375  df-trls 27468  df-eupth 27971
This theorem is referenced by: (None)
  Copyright terms: Public domain W3C validator