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

Theorem umgr3cyclex 29127
Description: If there are three (different) vertices in a multigraph which are mutually connected by edges, there is a 3-cycle in the graph containing one of these vertices. (Contributed by Alexander van der Vekens, 17-Nov-2017.) (Revised by AV, 12-Feb-2021.)
Hypotheses
Ref Expression
uhgr3cyclex.v 𝑉 = (Vtx‘𝐺)
uhgr3cyclex.e 𝐸 = (Edg‘𝐺)
Assertion
Ref Expression
umgr3cyclex ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ∃𝑓𝑝(𝑓(Cycles‘𝐺)𝑝 ∧ (♯‘𝑓) = 3 ∧ (𝑝‘0) = 𝐴))
Distinct variable groups:   𝐴,𝑓,𝑝   𝐵,𝑓,𝑝   𝐶,𝑓,𝑝   𝑓,𝐺,𝑝
Allowed substitution hints:   𝐸(𝑓,𝑝)   𝑉(𝑓,𝑝)

Proof of Theorem umgr3cyclex
StepHypRef Expression
1 umgruhgr 28055 . . 3 (𝐺 ∈ UMGraph → 𝐺 ∈ UHGraph)
213ad2ant1 1133 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐺 ∈ UHGraph)
3 simp2 1137 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝑉𝐵𝑉𝐶𝑉))
4 uhgr3cyclex.e . . . . . 6 𝐸 = (Edg‘𝐺)
54umgredgne 28096 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐴, 𝐵} ∈ 𝐸) → 𝐴𝐵)
653ad2antr1 1188 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐴𝐵)
7 prcom 4693 . . . . . . . 8 {𝐶, 𝐴} = {𝐴, 𝐶}
87eleq1i 2828 . . . . . . 7 ({𝐶, 𝐴} ∈ 𝐸 ↔ {𝐴, 𝐶} ∈ 𝐸)
98biimpi 215 . . . . . 6 ({𝐶, 𝐴} ∈ 𝐸 → {𝐴, 𝐶} ∈ 𝐸)
1093ad2ant3 1135 . . . . 5 (({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸) → {𝐴, 𝐶} ∈ 𝐸)
114umgredgne 28096 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐴, 𝐶} ∈ 𝐸) → 𝐴𝐶)
1210, 11sylan2 593 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐴𝐶)
13 simp2 1137 . . . . 5 (({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸) → {𝐵, 𝐶} ∈ 𝐸)
144umgredgne 28096 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐵, 𝐶} ∈ 𝐸) → 𝐵𝐶)
1513, 14sylan2 593 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐵𝐶)
166, 12, 153jca 1128 . . 3 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝐵𝐴𝐶𝐵𝐶))
17163adant2 1131 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝐵𝐴𝐶𝐵𝐶))
18 simp3 1138 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸))
19 uhgr3cyclex.v . . 3 𝑉 = (Vtx‘𝐺)
2019, 4uhgr3cyclex 29126 . 2 ((𝐺 ∈ UHGraph ∧ ((𝐴𝑉𝐵𝑉𝐶𝑉) ∧ (𝐴𝐵𝐴𝐶𝐵𝐶)) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ∃𝑓𝑝(𝑓(Cycles‘𝐺)𝑝 ∧ (♯‘𝑓) = 3 ∧ (𝑝‘0) = 𝐴))
212, 3, 17, 18, 20syl121anc 1375 1 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ∃𝑓𝑝(𝑓(Cycles‘𝐺)𝑝 ∧ (♯‘𝑓) = 3 ∧ (𝑝‘0) = 𝐴))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 396  w3a 1087   = wceq 1541  wex 1781  wcel 2106  wne 2943  {cpr 4588   class class class wbr 5105  cfv 6496  0cc0 11051  3c3 12209  chash 14230  Vtxcvtx 27947  Edgcedg 27998  UHGraphcuhgr 28007  UMGraphcumgr 28032  Cyclesccycls 28733
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1797  ax-4 1811  ax-5 1913  ax-6 1971  ax-7 2011  ax-8 2108  ax-9 2116  ax-10 2137  ax-11 2154  ax-12 2171  ax-ext 2707  ax-rep 5242  ax-sep 5256  ax-nul 5263  ax-pow 5320  ax-pr 5384  ax-un 7672  ax-cnex 11107  ax-resscn 11108  ax-1cn 11109  ax-icn 11110  ax-addcl 11111  ax-addrcl 11112  ax-mulcl 11113  ax-mulrcl 11114  ax-mulcom 11115  ax-addass 11116  ax-mulass 11117  ax-distr 11118  ax-i2m1 11119  ax-1ne0 11120  ax-1rid 11121  ax-rnegex 11122  ax-rrecex 11123  ax-cnre 11124  ax-pre-lttri 11125  ax-pre-lttrn 11126  ax-pre-ltadd 11127  ax-pre-mulgt0 11128
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 846  df-ifp 1062  df-3or 1088  df-3an 1089  df-tru 1544  df-fal 1554  df-ex 1782  df-nf 1786  df-sb 2068  df-mo 2538  df-eu 2567  df-clab 2714  df-cleq 2728  df-clel 2814  df-nfc 2889  df-ne 2944  df-nel 3050  df-ral 3065  df-rex 3074  df-reu 3354  df-rab 3408  df-v 3447  df-sbc 3740  df-csb 3856  df-dif 3913  df-un 3915  df-in 3917  df-ss 3927  df-pss 3929  df-nul 4283  df-if 4487  df-pw 4562  df-sn 4587  df-pr 4589  df-tp 4591  df-op 4593  df-uni 4866  df-int 4908  df-iun 4956  df-br 5106  df-opab 5168  df-mpt 5189  df-tr 5223  df-id 5531  df-eprel 5537  df-po 5545  df-so 5546  df-fr 5588  df-we 5590  df-xp 5639  df-rel 5640  df-cnv 5641  df-co 5642  df-dm 5643  df-rn 5644  df-res 5645  df-ima 5646  df-pred 6253  df-ord 6320  df-on 6321  df-lim 6322  df-suc 6323  df-iota 6448  df-fun 6498  df-fn 6499  df-f 6500  df-f1 6501  df-fo 6502  df-f1o 6503  df-fv 6504  df-riota 7313  df-ov 7360  df-oprab 7361  df-mpo 7362  df-om 7803  df-1st 7921  df-2nd 7922  df-frecs 8212  df-wrecs 8243  df-recs 8317  df-rdg 8356  df-1o 8412  df-oadd 8416  df-er 8648  df-map 8767  df-en 8884  df-dom 8885  df-sdom 8886  df-fin 8887  df-dju 9837  df-card 9875  df-pnf 11191  df-mnf 11192  df-xr 11193  df-ltxr 11194  df-le 11195  df-sub 11387  df-neg 11388  df-nn 12154  df-2 12216  df-3 12217  df-4 12218  df-n0 12414  df-z 12500  df-uz 12764  df-fz 13425  df-fzo 13568  df-hash 14231  df-word 14403  df-concat 14459  df-s1 14484  df-s2 14737  df-s3 14738  df-s4 14739  df-edg 27999  df-uhgr 28009  df-upgr 28033  df-umgr 28034  df-wlks 28547  df-trls 28640  df-pths 28664  df-cycls 28735
This theorem is referenced by:  umgr3v3e3cycl  29128  3cyclfrgr  29232  cusgr3cyclex  33730
  Copyright terms: Public domain W3C validator