Metamath Proof Explorer


Theorem 0cycl

Description: A pair of an empty set (of edges) and a second set (of vertices) is a cycle if and only if the second set contains exactly one vertex (in an undirected graph). (Contributed by Alexander van der Vekens, 30-Oct-2017) (Revised by AV, 31-Jan-2021) (Revised by AV, 30-Oct-2021)

Ref Expression
Assertion 0cycl ( 𝐺𝑊 → ( ∅ ( Cycles ‘ 𝐺 ) 𝑃𝑃 : ( 0 ... 0 ) ⟶ ( Vtx ‘ 𝐺 ) ) )

Proof

Step Hyp Ref Expression
1 eqid ( Vtx ‘ 𝐺 ) = ( Vtx ‘ 𝐺 )
2 1 0pth ( 𝐺𝑊 → ( ∅ ( Paths ‘ 𝐺 ) 𝑃𝑃 : ( 0 ... 0 ) ⟶ ( Vtx ‘ 𝐺 ) ) )
3 2 anbi1d ( 𝐺𝑊 → ( ( ∅ ( Paths ‘ 𝐺 ) 𝑃 ∧ ( 𝑃 ‘ 0 ) = ( 𝑃 ‘ ( ♯ ‘ ∅ ) ) ) ↔ ( 𝑃 : ( 0 ... 0 ) ⟶ ( Vtx ‘ 𝐺 ) ∧ ( 𝑃 ‘ 0 ) = ( 𝑃 ‘ ( ♯ ‘ ∅ ) ) ) ) )
4 iscycl ( ∅ ( Cycles ‘ 𝐺 ) 𝑃 ↔ ( ∅ ( Paths ‘ 𝐺 ) 𝑃 ∧ ( 𝑃 ‘ 0 ) = ( 𝑃 ‘ ( ♯ ‘ ∅ ) ) ) )
5 hash0 ( ♯ ‘ ∅ ) = 0
6 5 eqcomi 0 = ( ♯ ‘ ∅ )
7 6 fveq2i ( 𝑃 ‘ 0 ) = ( 𝑃 ‘ ( ♯ ‘ ∅ ) )
8 7 biantru ( 𝑃 : ( 0 ... 0 ) ⟶ ( Vtx ‘ 𝐺 ) ↔ ( 𝑃 : ( 0 ... 0 ) ⟶ ( Vtx ‘ 𝐺 ) ∧ ( 𝑃 ‘ 0 ) = ( 𝑃 ‘ ( ♯ ‘ ∅ ) ) ) )
9 3 4 8 3bitr4g ( 𝐺𝑊 → ( ∅ ( Cycles ‘ 𝐺 ) 𝑃𝑃 : ( 0 ... 0 ) ⟶ ( Vtx ‘ 𝐺 ) ) )