| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > df-trls | Structured version Visualization version GIF version | ||
| Description: Define the set of all
Trails (in an undirected graph).
According to Wikipedia ("Path (graph theory)", https://en.wikipedia.org/wiki/Path_(graph_theory), 3-Oct-2017): "A trail is a walk in which all edges are distinct. According to Bollobas: "... walk is called a trail if all its edges are distinct.", see Definition of [Bollobas] p. 5. Therefore, a trail can be represented by an injective mapping f from { 1 , ... , n } and a mapping p from { 0 , ... , n }, where f enumerates the (indices of the) different edges, and p enumerates the vertices. So the trail is also 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.) (Revised by AV, 28-Dec-2020.) |
| Ref | Expression |
|---|---|
| df-trls | ⊢ Trails = (𝑔 ∈ V ↦ {〈𝑓, 𝑝〉 ∣ (𝑓(Walks‘𝑔)𝑝 ∧ Fun ◡𝑓)}) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | ctrls 29708 | . 2 class Trails | |
| 2 | vg | . . 3 setvar 𝑔 | |
| 3 | cvv 3480 | . . 3 class V | |
| 4 | vf | . . . . . . 7 setvar 𝑓 | |
| 5 | 4 | cv 1539 | . . . . . 6 class 𝑓 |
| 6 | vp | . . . . . . 7 setvar 𝑝 | |
| 7 | 6 | cv 1539 | . . . . . 6 class 𝑝 |
| 8 | 2 | cv 1539 | . . . . . . 7 class 𝑔 |
| 9 | cwlks 29614 | . . . . . . 7 class Walks | |
| 10 | 8, 9 | cfv 6561 | . . . . . 6 class (Walks‘𝑔) |
| 11 | 5, 7, 10 | wbr 5143 | . . . . 5 wff 𝑓(Walks‘𝑔)𝑝 |
| 12 | 5 | ccnv 5684 | . . . . . 6 class ◡𝑓 |
| 13 | 12 | wfun 6555 | . . . . 5 wff Fun ◡𝑓 |
| 14 | 11, 13 | wa 395 | . . . 4 wff (𝑓(Walks‘𝑔)𝑝 ∧ Fun ◡𝑓) |
| 15 | 14, 4, 6 | copab 5205 | . . 3 class {〈𝑓, 𝑝〉 ∣ (𝑓(Walks‘𝑔)𝑝 ∧ Fun ◡𝑓)} |
| 16 | 2, 3, 15 | cmpt 5225 | . 2 class (𝑔 ∈ V ↦ {〈𝑓, 𝑝〉 ∣ (𝑓(Walks‘𝑔)𝑝 ∧ Fun ◡𝑓)}) |
| 17 | 1, 16 | wceq 1540 | 1 wff Trails = (𝑔 ∈ V ↦ {〈𝑓, 𝑝〉 ∣ (𝑓(Walks‘𝑔)𝑝 ∧ Fun ◡𝑓)}) |
| Colors of variables: wff setvar class |
| This definition is referenced by: reltrls 29712 trlsfval 29713 |
| Copyright terms: Public domain | W3C validator |