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

Theorem umgr3cyclex 27702
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 26582 . . 3 (𝐺 ∈ UMGraph → 𝐺 ∈ UHGraph)
213ad2ant1 1113 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐺 ∈ UHGraph)
3 simp2 1117 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝑉𝐵𝑉𝐶𝑉))
4 uhgr3cyclex.e . . . . . 6 𝐸 = (Edg‘𝐺)
54umgredgne 26623 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐴, 𝐵} ∈ 𝐸) → 𝐴𝐵)
653ad2antr1 1168 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐴𝐵)
7 prcom 4536 . . . . . . . 8 {𝐶, 𝐴} = {𝐴, 𝐶}
87eleq1i 2850 . . . . . . 7 ({𝐶, 𝐴} ∈ 𝐸 ↔ {𝐴, 𝐶} ∈ 𝐸)
98biimpi 208 . . . . . 6 ({𝐶, 𝐴} ∈ 𝐸 → {𝐴, 𝐶} ∈ 𝐸)
1093ad2ant3 1115 . . . . 5 (({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸) → {𝐴, 𝐶} ∈ 𝐸)
114umgredgne 26623 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐴, 𝐶} ∈ 𝐸) → 𝐴𝐶)
1210, 11sylan2 583 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐴𝐶)
13 simp2 1117 . . . . 5 (({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸) → {𝐵, 𝐶} ∈ 𝐸)
144umgredgne 26623 . . . . 5 ((𝐺 ∈ UMGraph ∧ {𝐵, 𝐶} ∈ 𝐸) → 𝐵𝐶)
1513, 14sylan2 583 . . . 4 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → 𝐵𝐶)
166, 12, 153jca 1108 . . 3 ((𝐺 ∈ UMGraph ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝐵𝐴𝐶𝐵𝐶))
17163adant2 1111 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → (𝐴𝐵𝐴𝐶𝐵𝐶))
18 simp3 1118 . 2 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸))
19 uhgr3cyclex.v . . 3 𝑉 = (Vtx‘𝐺)
2019, 4uhgr3cyclex 27701 . 2 ((𝐺 ∈ UHGraph ∧ ((𝐴𝑉𝐵𝑉𝐶𝑉) ∧ (𝐴𝐵𝐴𝐶𝐵𝐶)) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ∃𝑓𝑝(𝑓(Cycles‘𝐺)𝑝 ∧ (♯‘𝑓) = 3 ∧ (𝑝‘0) = 𝐴))
212, 3, 17, 18, 20syl121anc 1355 1 ((𝐺 ∈ UMGraph ∧ (𝐴𝑉𝐵𝑉𝐶𝑉) ∧ ({𝐴, 𝐵} ∈ 𝐸 ∧ {𝐵, 𝐶} ∈ 𝐸 ∧ {𝐶, 𝐴} ∈ 𝐸)) → ∃𝑓𝑝(𝑓(Cycles‘𝐺)𝑝 ∧ (♯‘𝑓) = 3 ∧ (𝑝‘0) = 𝐴))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 387  w3a 1068   = wceq 1507  wex 1742  wcel 2048  wne 2961  {cpr 4437   class class class wbr 4923  cfv 6182  0cc0 10327  3c3 11489  chash 13498  Vtxcvtx 26474  Edgcedg 26525  UHGraphcuhgr 26534  UMGraphcumgr 26559  Cyclesccycls 27264
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1758  ax-4 1772  ax-5 1869  ax-6 1928  ax-7 1964  ax-8 2050  ax-9 2057  ax-10 2077  ax-11 2091  ax-12 2104  ax-13 2299  ax-ext 2745  ax-rep 5043  ax-sep 5054  ax-nul 5061  ax-pow 5113  ax-pr 5180  ax-un 7273  ax-cnex 10383  ax-resscn 10384  ax-1cn 10385  ax-icn 10386  ax-addcl 10387  ax-addrcl 10388  ax-mulcl 10389  ax-mulrcl 10390  ax-mulcom 10391  ax-addass 10392  ax-mulass 10393  ax-distr 10394  ax-i2m1 10395  ax-1ne0 10396  ax-1rid 10397  ax-rnegex 10398  ax-rrecex 10399  ax-cnre 10400  ax-pre-lttri 10401  ax-pre-lttrn 10402  ax-pre-ltadd 10403  ax-pre-mulgt0 10404
This theorem depends on definitions:  df-bi 199  df-an 388  df-or 834  df-ifp 1044  df-3or 1069  df-3an 1070  df-tru 1510  df-ex 1743  df-nf 1747  df-sb 2014  df-mo 2544  df-eu 2580  df-clab 2754  df-cleq 2765  df-clel 2840  df-nfc 2912  df-ne 2962  df-nel 3068  df-ral 3087  df-rex 3088  df-reu 3089  df-rmo 3090  df-rab 3091  df-v 3411  df-sbc 3678  df-csb 3783  df-dif 3828  df-un 3830  df-in 3832  df-ss 3839  df-pss 3841  df-nul 4174  df-if 4345  df-pw 4418  df-sn 4436  df-pr 4438  df-tp 4440  df-op 4442  df-uni 4707  df-int 4744  df-iun 4788  df-br 4924  df-opab 4986  df-mpt 5003  df-tr 5025  df-id 5305  df-eprel 5310  df-po 5319  df-so 5320  df-fr 5359  df-we 5361  df-xp 5406  df-rel 5407  df-cnv 5408  df-co 5409  df-dm 5410  df-rn 5411  df-res 5412  df-ima 5413  df-pred 5980  df-ord 6026  df-on 6027  df-lim 6028  df-suc 6029  df-iota 6146  df-fun 6184  df-fn 6185  df-f 6186  df-f1 6187  df-fo 6188  df-f1o 6189  df-fv 6190  df-riota 6931  df-ov 6973  df-oprab 6974  df-mpo 6975  df-om 7391  df-1st 7494  df-2nd 7495  df-wrecs 7743  df-recs 7805  df-rdg 7843  df-1o 7897  df-oadd 7901  df-er 8081  df-map 8200  df-en 8299  df-dom 8300  df-sdom 8301  df-fin 8302  df-dju 9116  df-card 9154  df-pnf 10468  df-mnf 10469  df-xr 10470  df-ltxr 10471  df-le 10472  df-sub 10664  df-neg 10665  df-nn 11432  df-2 11496  df-3 11497  df-4 11498  df-n0 11701  df-z 11787  df-uz 12052  df-fz 12702  df-fzo 12843  df-hash 13499  df-word 13663  df-concat 13724  df-s1 13749  df-s2 14062  df-s3 14063  df-s4 14064  df-edg 26526  df-uhgr 26536  df-upgr 26560  df-umgr 26561  df-wlks 27074  df-trls 27170  df-pths 27195  df-cycls 27266
This theorem is referenced by:  umgr3v3e3cycl  27703  3cyclfrgr  27812
  Copyright terms: Public domain W3C validator