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

Definition df-wlk 26036
Description: Define the set of all Walks (in an undirected graph).

According to Wikipedia ("Path (graph theory)", https://en.wikipedia.org/wiki/Path_(graph_theory), 3-Oct-2017): "A walk of length k in a graph is an alternating sequence of vertices and edges, v0 , e0 , v1 , e1 , v2 , ... , v(k-1) , e(k-1) , v(k) which begins and ends with vertices. If the graph is undirected, then the endpoints of e(i) are v(i) and v(i+1)."

According to Bollobas: " A walk W in a graph is an alternating sequence of vertices and edges x0 , e1 , x1 , e2 , ... , e(l) , x(l) where e(i) = x(i-1)x(i), 0<i<=l.", see Definition of [Bollobas] p. 4.

Therefore, a walk can be represented by two mappings f from { 1 , ... , n } and p from { 0 , ... , n }, where f enumerates the (indices of the) edges, and p enumerates the vertices. So the walk is represented by the following sequence: p(0) e(f(1)) p(1) e(f(2)) ... p(n-1) e(f(n)) p(n). (Contributed by Alexander van der Vekens and Mario Carneiro, 4-Oct-2017.)

Assertion
Ref Expression
df-wlk Walks = (𝑣 ∈ V, 𝑒 ∈ V ↦ {⟨𝑓, 𝑝⟩ ∣ (𝑓 ∈ Word dom 𝑒𝑝:(0...(#‘𝑓))⟶𝑣 ∧ ∀𝑘 ∈ (0..^(#‘𝑓))(𝑒‘(𝑓𝑘)) = {(𝑝𝑘), (𝑝‘(𝑘 + 1))})})
Distinct variable group:   𝑣,𝑒,𝑘,𝑓,𝑝

Detailed syntax breakdown of Definition df-wlk
StepHypRef Expression
1 cwalk 26026 . 2 class Walks
2 vv . . 3 setvar 𝑣
3 ve . . 3 setvar 𝑒
4 cvv 3173 . . 3 class V
5 vf . . . . . . 7 setvar 𝑓
65cv 1474 . . . . . 6 class 𝑓
73cv 1474 . . . . . . . 8 class 𝑒
87cdm 5038 . . . . . . 7 class dom 𝑒
98cword 13146 . . . . . 6 class Word dom 𝑒
106, 9wcel 1977 . . . . 5 wff 𝑓 ∈ Word dom 𝑒
11 cc0 9815 . . . . . . 7 class 0
12 chash 12979 . . . . . . . 8 class #
136, 12cfv 5804 . . . . . . 7 class (#‘𝑓)
14 cfz 12197 . . . . . . 7 class ...
1511, 13, 14co 6549 . . . . . 6 class (0...(#‘𝑓))
162cv 1474 . . . . . 6 class 𝑣
17 vp . . . . . . 7 setvar 𝑝
1817cv 1474 . . . . . 6 class 𝑝
1915, 16, 18wf 5800 . . . . 5 wff 𝑝:(0...(#‘𝑓))⟶𝑣
20 vk . . . . . . . . . 10 setvar 𝑘
2120cv 1474 . . . . . . . . 9 class 𝑘
2221, 6cfv 5804 . . . . . . . 8 class (𝑓𝑘)
2322, 7cfv 5804 . . . . . . 7 class (𝑒‘(𝑓𝑘))
2421, 18cfv 5804 . . . . . . . 8 class (𝑝𝑘)
25 c1 9816 . . . . . . . . . 10 class 1
26 caddc 9818 . . . . . . . . . 10 class +
2721, 25, 26co 6549 . . . . . . . . 9 class (𝑘 + 1)
2827, 18cfv 5804 . . . . . . . 8 class (𝑝‘(𝑘 + 1))
2924, 28cpr 4127 . . . . . . 7 class {(𝑝𝑘), (𝑝‘(𝑘 + 1))}
3023, 29wceq 1475 . . . . . 6 wff (𝑒‘(𝑓𝑘)) = {(𝑝𝑘), (𝑝‘(𝑘 + 1))}
31 cfzo 12334 . . . . . . 7 class ..^
3211, 13, 31co 6549 . . . . . 6 class (0..^(#‘𝑓))
3330, 20, 32wral 2896 . . . . 5 wff 𝑘 ∈ (0..^(#‘𝑓))(𝑒‘(𝑓𝑘)) = {(𝑝𝑘), (𝑝‘(𝑘 + 1))}
3410, 19, 33w3a 1031 . . . 4 wff (𝑓 ∈ Word dom 𝑒𝑝:(0...(#‘𝑓))⟶𝑣 ∧ ∀𝑘 ∈ (0..^(#‘𝑓))(𝑒‘(𝑓𝑘)) = {(𝑝𝑘), (𝑝‘(𝑘 + 1))})
3534, 5, 17copab 4642 . . 3 class {⟨𝑓, 𝑝⟩ ∣ (𝑓 ∈ Word dom 𝑒𝑝:(0...(#‘𝑓))⟶𝑣 ∧ ∀𝑘 ∈ (0..^(#‘𝑓))(𝑒‘(𝑓𝑘)) = {(𝑝𝑘), (𝑝‘(𝑘 + 1))})}
362, 3, 4, 4, 35cmpt2 6551 . 2 class (𝑣 ∈ V, 𝑒 ∈ V ↦ {⟨𝑓, 𝑝⟩ ∣ (𝑓 ∈ Word dom 𝑒𝑝:(0...(#‘𝑓))⟶𝑣 ∧ ∀𝑘 ∈ (0..^(#‘𝑓))(𝑒‘(𝑓𝑘)) = {(𝑝𝑘), (𝑝‘(𝑘 + 1))})})
371, 36wceq 1475 1 wff Walks = (𝑣 ∈ V, 𝑒 ∈ V ↦ {⟨𝑓, 𝑝⟩ ∣ (𝑓 ∈ Word dom 𝑒𝑝:(0...(#‘𝑓))⟶𝑣 ∧ ∀𝑘 ∈ (0..^(#‘𝑓))(𝑒‘(𝑓𝑘)) = {(𝑝𝑘), (𝑝‘(𝑘 + 1))})})
Colors of variables: wff setvar class
This definition is referenced by:  relwlk  26046  wlks  26047  2mwlk  26049  wlkbprop  26051  wlkcompim  26054  wlkelwrd  26058  wlkdvspth  26138
  Copyright terms: Public domain W3C validator