| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > df-clwlks | Structured version Visualization version GIF version | ||
| Description: Define the set of all
closed walks (in an undirected graph).
According to definition 4 in [Huneke] p. 2: "A walk of length n on (a graph) G is an ordered sequence v0 , v1 , ... v(n) of vertices such that v(i) and v(i+1) are neighbors (i.e are connected by an edge). We say the walk is closed if v(n) = v0". According to the definition of a walk as two mappings f from { 0 , ... , ( n - 1 ) } and p from { 0 , ... , n }, where f enumerates the (indices of the) edges, and p enumerates the vertices, a closed walk is represented by the following sequence: p(0) e(f(0)) p(1) e(f(1)) ... p(n-1) e(f(n-1)) p(n)=p(0). Notice that by this definition, a single vertex is a closed walk of length 0, see also 0clwlk 26991! (Contributed by Alexander van der Vekens, 12-Mar-2018.) (Revised by AV, 16-Feb-2021.) |
| Ref | Expression |
|---|---|
| df-clwlks | ⊢ ClWalks = (𝑔 ∈ V ↦ {〈𝑓, 𝑝〉 ∣ (𝑓(Walks‘𝑔)𝑝 ∧ (𝑝‘0) = (𝑝‘(#‘𝑓)))}) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | cclwlks 26666 | . 2 class ClWalks | |
| 2 | vg | . . 3 setvar 𝑔 | |
| 3 | cvv 3200 | . . 3 class V | |
| 4 | vf | . . . . . . 7 setvar 𝑓 | |
| 5 | 4 | cv 1482 | . . . . . 6 class 𝑓 |
| 6 | vp | . . . . . . 7 setvar 𝑝 | |
| 7 | 6 | cv 1482 | . . . . . 6 class 𝑝 |
| 8 | 2 | cv 1482 | . . . . . . 7 class 𝑔 |
| 9 | cwlks 26492 | . . . . . . 7 class Walks | |
| 10 | 8, 9 | cfv 5888 | . . . . . 6 class (Walks‘𝑔) |
| 11 | 5, 7, 10 | wbr 4653 | . . . . 5 wff 𝑓(Walks‘𝑔)𝑝 |
| 12 | cc0 9936 | . . . . . . 7 class 0 | |
| 13 | 12, 7 | cfv 5888 | . . . . . 6 class (𝑝‘0) |
| 14 | chash 13117 | . . . . . . . 8 class # | |
| 15 | 5, 14 | cfv 5888 | . . . . . . 7 class (#‘𝑓) |
| 16 | 15, 7 | cfv 5888 | . . . . . 6 class (𝑝‘(#‘𝑓)) |
| 17 | 13, 16 | wceq 1483 | . . . . 5 wff (𝑝‘0) = (𝑝‘(#‘𝑓)) |
| 18 | 11, 17 | wa 384 | . . . 4 wff (𝑓(Walks‘𝑔)𝑝 ∧ (𝑝‘0) = (𝑝‘(#‘𝑓))) |
| 19 | 18, 4, 6 | copab 4712 | . . 3 class {〈𝑓, 𝑝〉 ∣ (𝑓(Walks‘𝑔)𝑝 ∧ (𝑝‘0) = (𝑝‘(#‘𝑓)))} |
| 20 | 2, 3, 19 | cmpt 4729 | . 2 class (𝑔 ∈ V ↦ {〈𝑓, 𝑝〉 ∣ (𝑓(Walks‘𝑔)𝑝 ∧ (𝑝‘0) = (𝑝‘(#‘𝑓)))}) |
| 21 | 1, 20 | wceq 1483 | 1 wff ClWalks = (𝑔 ∈ V ↦ {〈𝑓, 𝑝〉 ∣ (𝑓(Walks‘𝑔)𝑝 ∧ (𝑝‘0) = (𝑝‘(#‘𝑓)))}) |
| Colors of variables: wff setvar class |
| This definition is referenced by: clwlks 26668 |
| Copyright terms: Public domain | W3C validator |