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

Theorem umgr3cyclex 27889
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 26816 . . 3 (𝐺 ∈ UMGraph → 𝐺 ∈ UHGraph)
213ad2ant1 1125 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐺 ∈ UHGraph)
3 simp2 1129 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝑉𝐵𝑉𝐶𝑉))
4 uhgr3cyclex.e . . . . . 6 𝐸 = (Edg‘𝐺)
54umgredgne 26857 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐴, 𝐵} ∈ 𝐸) → 𝐴𝐵)
653ad2antr1 1180 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐴𝐵)
7 prcom 4660 . . . . . . . 8 {𝐶, 𝐴} = {𝐴, 𝐶}
87eleq1i 2900 . . . . . . 7 ({𝐶, 𝐴} ∈ 𝐸 ↔ {𝐴, 𝐶} ∈ 𝐸)
98biimpi 217 . . . . . 6 ({𝐶, 𝐴} ∈ 𝐸 → {𝐴, 𝐶} ∈ 𝐸)
1093ad2ant3 1127 . . . . 5 (({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸) → {𝐴, 𝐶} ∈ 𝐸)
114umgredgne 26857 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐴, 𝐶} ∈ 𝐸) → 𝐴𝐶)
1210, 11sylan2 592 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐴𝐶)
13 simp2 1129 . . . . 5 (({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸) → {𝐵, 𝐶} ∈ 𝐸)
144umgredgne 26857 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐵, 𝐶} ∈ 𝐸) → 𝐵𝐶)
1513, 14sylan2 592 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐵𝐶)
166, 12, 153jca 1120 . . 3 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝐵𝐴𝐶𝐵𝐶))
17163adant2 1123 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝐵𝐴𝐶𝐵𝐶))
18 simp3 1130 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸))
19 uhgr3cyclex.v . . 3 𝑉 = (Vtx‘𝐺)
2019, 4uhgr3cyclex 27888 . 2 ((𝐺 ∈ UHGraph ∧ ((𝐴𝑉𝐵𝑉𝐶𝑉) ∧ (𝐴𝐵𝐴𝐶𝐵𝐶)) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ∃𝑓𝑝(𝑓(Cycles‘𝐺)𝑝 ∧ (♯‘𝑓) = 3 ∧ (𝑝‘0) = 𝐴))
212, 3, 17, 18, 20syl121anc 1367 1 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ∃𝑓𝑝(𝑓(Cycles‘𝐺)𝑝 ∧ (♯‘𝑓) = 3 ∧ (𝑝‘0) = 𝐴))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 396  w3a 1079   = wceq 1528  wex 1771  wcel 2105  wne 3013  {cpr 4559   class class class wbr 5057  cfv 6348  0cc0 10525  3c3 11681  chash 13678  Vtxcvtx 26708  Edgcedg 26759  UHGraphcuhgr 26768  UMGraphcumgr 26793  Cyclesccycls 27493
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1787  ax-4 1801  ax-5 1902  ax-6 1961  ax-7 2006  ax-8 2107  ax-9 2115  ax-10 2136  ax-11 2151  ax-12 2167  ax-ext 2790  ax-rep 5181  ax-sep 5194  ax-nul 5201  ax-pow 5257  ax-pr 5320  ax-un 7450  ax-cnex 10581  ax-resscn 10582  ax-1cn 10583  ax-icn 10584  ax-addcl 10585  ax-addrcl 10586  ax-mulcl 10587  ax-mulrcl 10588  ax-mulcom 10589  ax-addass 10590  ax-mulass 10591  ax-distr 10592  ax-i2m1 10593  ax-1ne0 10594  ax-1rid 10595  ax-rnegex 10596  ax-rrecex 10597  ax-cnre 10598  ax-pre-lttri 10599  ax-pre-lttrn 10600  ax-pre-ltadd 10601  ax-pre-mulgt0 10602
This theorem depends on definitions:  df-bi 208  df-an 397  df-or 842  df-ifp 1055  df-3or 1080  df-3an 1081  df-tru 1531  df-ex 1772  df-nf 1776  df-sb 2061  df-mo 2615  df-eu 2647  df-clab 2797  df-cleq 2811  df-clel 2890  df-nfc 2960  df-ne 3014  df-nel 3121  df-ral 3140  df-rex 3141  df-reu 3142  df-rmo 3143  df-rab 3144  df-v 3494  df-sbc 3770  df-csb 3881  df-dif 3936  df-un 3938  df-in 3940  df-ss 3949  df-pss 3951  df-nul 4289  df-if 4464  df-pw 4537  df-sn 4558  df-pr 4560  df-tp 4562  df-op 4564  df-uni 4831  df-int 4868  df-iun 4912  df-br 5058  df-opab 5120  df-mpt 5138  df-tr 5164  df-id 5453  df-eprel 5458  df-po 5467  df-so 5468  df-fr 5507  df-we 5509  df-xp 5554  df-rel 5555  df-cnv 5556  df-co 5557  df-dm 5558  df-rn 5559  df-res 5560  df-ima 5561  df-pred 6141  df-ord 6187  df-on 6188  df-lim 6189  df-suc 6190  df-iota 6307  df-fun 6350  df-fn 6351  df-f 6352  df-f1 6353  df-fo 6354  df-f1o 6355  df-fv 6356  df-riota 7103  df-ov 7148  df-oprab 7149  df-mpo 7150  df-om 7570  df-1st 7678  df-2nd 7679  df-wrecs 7936  df-recs 7997  df-rdg 8035  df-1o 8091  df-oadd 8095  df-er 8278  df-map 8397  df-en 8498  df-dom 8499  df-sdom 8500  df-fin 8501  df-dju 9318  df-card 9356  df-pnf 10665  df-mnf 10666  df-xr 10667  df-ltxr 10668  df-le 10669  df-sub 10860  df-neg 10861  df-nn 11627  df-2 11688  df-3 11689  df-4 11690  df-n0 11886  df-z 11970  df-uz 12232  df-fz 12881  df-fzo 13022  df-hash 13679  df-word 13850  df-concat 13911  df-s1 13938  df-s2 14198  df-s3 14199  df-s4 14200  df-edg 26760  df-uhgr 26770  df-upgr 26794  df-umgr 26795  df-wlks 27308  df-trls 27401  df-pths 27424  df-cycls 27495
This theorem is referenced by:  umgr3v3e3cycl  27890  3cyclfrgr  27994  cusgr3cyclex  32280
  Copyright terms: Public domain W3C validator