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

Theorem umgr3cyclex 30116
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 29040 . . 3 (𝐺 ∈ UMGraph → 𝐺 ∈ UHGraph)
213ad2ant1 1130 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐺 ∈ UHGraph)
3 simp2 1134 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝑉𝐵𝑉𝐶𝑉))
4 uhgr3cyclex.e . . . . . 6 𝐸 = (Edg‘𝐺)
54umgredgne 29081 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐴, 𝐵} ∈ 𝐸) → 𝐴𝐵)
653ad2antr1 1185 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐴𝐵)
7 prcom 4741 . . . . . . . 8 {𝐶, 𝐴} = {𝐴, 𝐶}
87eleq1i 2817 . . . . . . 7 ({𝐶, 𝐴} ∈ 𝐸 ↔ {𝐴, 𝐶} ∈ 𝐸)
98biimpi 215 . . . . . 6 ({𝐶, 𝐴} ∈ 𝐸 → {𝐴, 𝐶} ∈ 𝐸)
1093ad2ant3 1132 . . . . 5 (({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸) → {𝐴, 𝐶} ∈ 𝐸)
114umgredgne 29081 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐴, 𝐶} ∈ 𝐸) → 𝐴𝐶)
1210, 11sylan2 591 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐴𝐶)
13 simp2 1134 . . . . 5 (({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸) → {𝐵, 𝐶} ∈ 𝐸)
144umgredgne 29081 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐵, 𝐶} ∈ 𝐸) → 𝐵𝐶)
1513, 14sylan2 591 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐵𝐶)
166, 12, 153jca 1125 . . 3 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝐵𝐴𝐶𝐵𝐶))
17163adant2 1128 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝐵𝐴𝐶𝐵𝐶))
18 simp3 1135 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸))
19 uhgr3cyclex.v . . 3 𝑉 = (Vtx‘𝐺)
2019, 4uhgr3cyclex 30115 . 2 ((𝐺 ∈ UHGraph ∧ ((𝐴𝑉𝐵𝑉𝐶𝑉) ∧ (𝐴𝐵𝐴𝐶𝐵𝐶)) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ∃𝑓𝑝(𝑓(Cycles‘𝐺)𝑝 ∧ (♯‘𝑓) = 3 ∧ (𝑝‘0) = 𝐴))
212, 3, 17, 18, 20syl121anc 1372 1 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ∃𝑓𝑝(𝑓(Cycles‘𝐺)𝑝 ∧ (♯‘𝑓) = 3 ∧ (𝑝‘0) = 𝐴))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 394  w3a 1084   = wceq 1534  wex 1774  wcel 2099  wne 2930  {cpr 4635   class class class wbr 5153  cfv 6554  0cc0 11158  3c3 12320  chash 14347  Vtxcvtx 28932  Edgcedg 28983  UHGraphcuhgr 28992  UMGraphcumgr 29017  Cyclesccycls 29722
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1790  ax-4 1804  ax-5 1906  ax-6 1964  ax-7 2004  ax-8 2101  ax-9 2109  ax-10 2130  ax-11 2147  ax-12 2167  ax-ext 2697  ax-rep 5290  ax-sep 5304  ax-nul 5311  ax-pow 5369  ax-pr 5433  ax-un 7746  ax-cnex 11214  ax-resscn 11215  ax-1cn 11216  ax-icn 11217  ax-addcl 11218  ax-addrcl 11219  ax-mulcl 11220  ax-mulrcl 11221  ax-mulcom 11222  ax-addass 11223  ax-mulass 11224  ax-distr 11225  ax-i2m1 11226  ax-1ne0 11227  ax-1rid 11228  ax-rnegex 11229  ax-rrecex 11230  ax-cnre 11231  ax-pre-lttri 11232  ax-pre-lttrn 11233  ax-pre-ltadd 11234  ax-pre-mulgt0 11235
This theorem depends on definitions:  df-bi 206  df-an 395  df-or 846  df-ifp 1061  df-3or 1085  df-3an 1086  df-tru 1537  df-fal 1547  df-ex 1775  df-nf 1779  df-sb 2061  df-mo 2529  df-eu 2558  df-clab 2704  df-cleq 2718  df-clel 2803  df-nfc 2878  df-ne 2931  df-nel 3037  df-ral 3052  df-rex 3061  df-reu 3365  df-rab 3420  df-v 3464  df-sbc 3777  df-csb 3893  df-dif 3950  df-un 3952  df-in 3954  df-ss 3964  df-pss 3967  df-nul 4326  df-if 4534  df-pw 4609  df-sn 4634  df-pr 4636  df-tp 4638  df-op 4640  df-uni 4914  df-int 4955  df-iun 5003  df-br 5154  df-opab 5216  df-mpt 5237  df-tr 5271  df-id 5580  df-eprel 5586  df-po 5594  df-so 5595  df-fr 5637  df-we 5639  df-xp 5688  df-rel 5689  df-cnv 5690  df-co 5691  df-dm 5692  df-rn 5693  df-res 5694  df-ima 5695  df-pred 6312  df-ord 6379  df-on 6380  df-lim 6381  df-suc 6382  df-iota 6506  df-fun 6556  df-fn 6557  df-f 6558  df-f1 6559  df-fo 6560  df-f1o 6561  df-fv 6562  df-riota 7380  df-ov 7427  df-oprab 7428  df-mpo 7429  df-om 7877  df-1st 8003  df-2nd 8004  df-frecs 8296  df-wrecs 8327  df-recs 8401  df-rdg 8440  df-1o 8496  df-oadd 8500  df-er 8734  df-map 8857  df-en 8975  df-dom 8976  df-sdom 8977  df-fin 8978  df-dju 9944  df-card 9982  df-pnf 11300  df-mnf 11301  df-xr 11302  df-ltxr 11303  df-le 11304  df-sub 11496  df-neg 11497  df-nn 12265  df-2 12327  df-3 12328  df-4 12329  df-n0 12525  df-z 12611  df-uz 12875  df-fz 13539  df-fzo 13682  df-hash 14348  df-word 14523  df-concat 14579  df-s1 14604  df-s2 14857  df-s3 14858  df-s4 14859  df-edg 28984  df-uhgr 28994  df-upgr 29018  df-umgr 29019  df-wlks 29536  df-trls 29629  df-pths 29653  df-cycls 29724
This theorem is referenced by:  umgr3v3e3cycl  30117  3cyclfrgr  30221  cusgr3cyclex  34964
  Copyright terms: Public domain W3C validator