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

Theorem grstructd 29067
Description: If any representation of a graph with vertices 𝑉 and edges 𝐸 has a certain property 𝜓, then any structure with base set 𝑉 and value 𝐸 in the slot for edge functions (which is such a representation of a graph with vertices 𝑉 and edges 𝐸) has this property. (Contributed by AV, 12-Oct-2020.) (Revised by AV, 9-Jun-2021.)
Hypotheses
Ref Expression
gropd.g (𝜑 → ∀𝑔(((Vtx‘𝑔) = 𝑉 ∧ (iEdg‘𝑔) = 𝐸) → 𝜓))
gropd.v (𝜑𝑉𝑈)
gropd.e (𝜑𝐸𝑊)
grstructd.s (𝜑𝑆𝑋)
grstructd.f (𝜑 → Fun (𝑆 ∖ {∅}))
grstructd.d (𝜑 → 2 ≤ (♯‘dom 𝑆))
grstructd.b (𝜑 → (Base‘𝑆) = 𝑉)
grstructd.e (𝜑 → (.ef‘𝑆) = 𝐸)
Assertion
Ref Expression
grstructd (𝜑[𝑆 / 𝑔]𝜓)
Distinct variable groups:   𝑔,𝐸   𝑔,𝑉   𝜑,𝑔   𝑆,𝑔
Allowed substitution hints:   𝜓(𝑔)   𝑈(𝑔)   𝑊(𝑔)   𝑋(𝑔)

Proof of Theorem grstructd
StepHypRef Expression
1 grstructd.s . 2 (𝜑𝑆𝑋)
2 gropd.g . 2 (𝜑 → ∀𝑔(((Vtx‘𝑔) = 𝑉 ∧ (iEdg‘𝑔) = 𝐸) → 𝜓))
3 grstructd.f . . . . 5 (𝜑 → Fun (𝑆 ∖ {∅}))
4 grstructd.d . . . . 5 (𝜑 → 2 ≤ (♯‘dom 𝑆))
5 funvtxdmge2val 29046 . . . . 5 ((Fun (𝑆 ∖ {∅}) ∧ 2 ≤ (♯‘dom 𝑆)) → (Vtx‘𝑆) = (Base‘𝑆))
63, 4, 5syl2anc 583 . . . 4 (𝜑 → (Vtx‘𝑆) = (Base‘𝑆))
7 grstructd.b . . . 4 (𝜑 → (Base‘𝑆) = 𝑉)
86, 7eqtrd 2780 . . 3 (𝜑 → (Vtx‘𝑆) = 𝑉)
9 funiedgdmge2val 29047 . . . . 5 ((Fun (𝑆 ∖ {∅}) ∧ 2 ≤ (♯‘dom 𝑆)) → (iEdg‘𝑆) = (.ef‘𝑆))
103, 4, 9syl2anc 583 . . . 4 (𝜑 → (iEdg‘𝑆) = (.ef‘𝑆))
11 grstructd.e . . . 4 (𝜑 → (.ef‘𝑆) = 𝐸)
1210, 11eqtrd 2780 . . 3 (𝜑 → (iEdg‘𝑆) = 𝐸)
138, 12jca 511 . 2 (𝜑 → ((Vtx‘𝑆) = 𝑉 ∧ (iEdg‘𝑆) = 𝐸))
14 nfcv 2908 . . 3 𝑔𝑆
15 nfv 1913 . . . 4 𝑔((Vtx‘𝑆) = 𝑉 ∧ (iEdg‘𝑆) = 𝐸)
16 nfsbc1v 3824 . . . 4 𝑔[𝑆 / 𝑔]𝜓
1715, 16nfim 1895 . . 3 𝑔(((Vtx‘𝑆) = 𝑉 ∧ (iEdg‘𝑆) = 𝐸) → [𝑆 / 𝑔]𝜓)
18 fveqeq2 6929 . . . . 5 (𝑔 = 𝑆 → ((Vtx‘𝑔) = 𝑉 ↔ (Vtx‘𝑆) = 𝑉))
19 fveqeq2 6929 . . . . 5 (𝑔 = 𝑆 → ((iEdg‘𝑔) = 𝐸 ↔ (iEdg‘𝑆) = 𝐸))
2018, 19anbi12d 631 . . . 4 (𝑔 = 𝑆 → (((Vtx‘𝑔) = 𝑉 ∧ (iEdg‘𝑔) = 𝐸) ↔ ((Vtx‘𝑆) = 𝑉 ∧ (iEdg‘𝑆) = 𝐸)))
21 sbceq1a 3815 . . . 4 (𝑔 = 𝑆 → (𝜓[𝑆 / 𝑔]𝜓))
2220, 21imbi12d 344 . . 3 (𝑔 = 𝑆 → ((((Vtx‘𝑔) = 𝑉 ∧ (iEdg‘𝑔) = 𝐸) → 𝜓) ↔ (((Vtx‘𝑆) = 𝑉 ∧ (iEdg‘𝑆) = 𝐸) → [𝑆 / 𝑔]𝜓)))
2314, 17, 22spcgf 3604 . 2 (𝑆𝑋 → (∀𝑔(((Vtx‘𝑔) = 𝑉 ∧ (iEdg‘𝑔) = 𝐸) → 𝜓) → (((Vtx‘𝑆) = 𝑉 ∧ (iEdg‘𝑆) = 𝐸) → [𝑆 / 𝑔]𝜓)))
241, 2, 13, 23syl3c 66 1 (𝜑[𝑆 / 𝑔]𝜓)
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 395  wal 1535   = wceq 1537  wcel 2108  [wsbc 3804  cdif 3973  c0 4352  {csn 4648   class class class wbr 5166  dom cdm 5700  Fun wfun 6567  cfv 6573  cle 11325  2c2 12348  chash 14379  Basecbs 17258  .efcedgf 29021  Vtxcvtx 29031  iEdgciedg 29032
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1793  ax-4 1807  ax-5 1909  ax-6 1967  ax-7 2007  ax-8 2110  ax-9 2118  ax-10 2141  ax-11 2158  ax-12 2178  ax-ext 2711  ax-sep 5317  ax-nul 5324  ax-pow 5383  ax-pr 5447  ax-un 7770  ax-cnex 11240  ax-resscn 11241  ax-1cn 11242  ax-icn 11243  ax-addcl 11244  ax-addrcl 11245  ax-mulcl 11246  ax-mulrcl 11247  ax-mulcom 11248  ax-addass 11249  ax-mulass 11250  ax-distr 11251  ax-i2m1 11252  ax-1ne0 11253  ax-1rid 11254  ax-rnegex 11255  ax-rrecex 11256  ax-cnre 11257  ax-pre-lttri 11258  ax-pre-lttrn 11259  ax-pre-ltadd 11260  ax-pre-mulgt0 11261
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 847  df-3or 1088  df-3an 1089  df-tru 1540  df-fal 1550  df-ex 1778  df-nf 1782  df-sb 2065  df-mo 2543  df-eu 2572  df-clab 2718  df-cleq 2732  df-clel 2819  df-nfc 2895  df-ne 2947  df-nel 3053  df-ral 3068  df-rex 3077  df-reu 3389  df-rab 3444  df-v 3490  df-sbc 3805  df-csb 3922  df-dif 3979  df-un 3981  df-in 3983  df-ss 3993  df-pss 3996  df-nul 4353  df-if 4549  df-pw 4624  df-sn 4649  df-pr 4651  df-op 4655  df-uni 4932  df-int 4971  df-iun 5017  df-br 5167  df-opab 5229  df-mpt 5250  df-tr 5284  df-id 5593  df-eprel 5599  df-po 5607  df-so 5608  df-fr 5652  df-we 5654  df-xp 5706  df-rel 5707  df-cnv 5708  df-co 5709  df-dm 5710  df-rn 5711  df-res 5712  df-ima 5713  df-pred 6332  df-ord 6398  df-on 6399  df-lim 6400  df-suc 6401  df-iota 6525  df-fun 6575  df-fn 6576  df-f 6577  df-f1 6578  df-fo 6579  df-f1o 6580  df-fv 6581  df-riota 7404  df-ov 7451  df-oprab 7452  df-mpo 7453  df-om 7904  df-1st 8030  df-2nd 8031  df-frecs 8322  df-wrecs 8353  df-recs 8427  df-rdg 8466  df-1o 8522  df-er 8763  df-en 9004  df-dom 9005  df-sdom 9006  df-fin 9007  df-card 10008  df-pnf 11326  df-mnf 11327  df-xr 11328  df-ltxr 11329  df-le 11330  df-sub 11522  df-neg 11523  df-nn 12294  df-2 12356  df-n0 12554  df-xnn0 12626  df-z 12640  df-uz 12904  df-fz 13568  df-hash 14380  df-vtx 29033  df-iedg 29034
This theorem is referenced by:  grstructeld  29069
  Copyright terms: Public domain W3C validator