Mapping walks between graphs #
Functions that map walks between different graphs.
Main definitions #
SimpleGraph.Walk.map: The map on walks induced by a graph homomorphismSimpleGraph.Walk.mapLe: Map a walk to a supergraphSimpleGraph.Walk.transfer: Map a walk to another graph that contains its edgesSimpleGraph.Walk.induce: Map a walk that's fully contained in a set of vertices to the subgraph induced by that setSimpleGraph.Walk.toDeleteEdges: Map a walk that avoids a set of edges to the subgraph with those edges deletedSimpleGraph.Walk.toDeleteEdge: Map a walk that avoids an edge to the subgraph with that edge deleted
Tags #
walks
Mapping walks #
def
SimpleGraph.Walk.map
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v : V}
:
Given a graph homomorphism, map walks to walks.
Equations
Instances For
@[simp]
theorem
SimpleGraph.Walk.map_nil
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u : V}
:
@[simp]
theorem
SimpleGraph.Walk.map_cons
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v : V}
(p : G.Walk u v)
{w : V}
(h : G.Adj w u)
:
@[simp]
theorem
SimpleGraph.Walk.map_copy
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v u' v' : V}
(p : G.Walk u v)
(hu : u = u')
(hv : v = v')
:
@[simp]
@[simp]
theorem
SimpleGraph.Walk.map_map
{V : Type u}
{V' : Type v}
{V'' : Type w}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
{G'' : SimpleGraph V''}
(f : G →g G')
(f' : G' →g G'')
{u v : V}
(p : G.Walk u v)
:
theorem
SimpleGraph.Walk.map_eq_of_eq
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
{u v : V}
(p : G.Walk u v)
{f : G →g G'}
(f' : G →g G')
(h : f = f')
:
Unlike categories, for graphs vertex equality is an important notion, so needing to be able to work with equality of graph homomorphisms is a necessary evil.
@[simp]
theorem
SimpleGraph.Walk.map_eq_nil_iff
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u : V}
{p : G.Walk u u}
:
@[simp]
theorem
SimpleGraph.Walk.length_map
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v : V}
(p : G.Walk u v)
:
theorem
SimpleGraph.Walk.map_append
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v w : V}
(p : G.Walk u v)
(q : G.Walk v w)
:
@[simp]
theorem
SimpleGraph.Walk.reverse_map
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v : V}
(p : G.Walk u v)
:
@[simp]
theorem
SimpleGraph.Walk.support_map
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v : V}
(p : G.Walk u v)
:
@[simp]
theorem
SimpleGraph.Walk.darts_map
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v : V}
(p : G.Walk u v)
:
@[simp]
theorem
SimpleGraph.Walk.edges_map
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v : V}
(p : G.Walk u v)
:
@[simp]
theorem
SimpleGraph.Walk.edgeSet_map
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v : V}
(p : G.Walk u v)
:
@[simp]
theorem
SimpleGraph.Walk.getVert_map
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
(f : G →g G')
{u v : V}
(p : G.Walk u v)
(n : ℕ)
:
theorem
SimpleGraph.Walk.map_injective_of_injective
{V : Type u}
{V' : Type v}
{G : SimpleGraph V}
{G' : SimpleGraph V'}
{f : G →g G'}
(hinj : Function.Injective ⇑f)
(u v : V)
:
@[reducible, inline]
abbrev
SimpleGraph.Walk.mapLe
{V : Type u}
{G G' : SimpleGraph V}
(h : G ≤ G')
{u v : V}
(p : G.Walk u v)
:
G'.Walk u v
The specialization of SimpleGraph.Walk.map for mapping walks to supergraphs.
Equations
Instances For
theorem
SimpleGraph.Walk.support_mapLe_eq_support
{V : Type u}
{G G' : SimpleGraph V}
(h : G ≤ G')
{u v : V}
(p : G.Walk u v)
:
theorem
SimpleGraph.Walk.edges_mapLe_eq_edges
{V : Type u}
{G G' : SimpleGraph V}
(h : G ≤ G')
{u v : V}
(p : G.Walk u v)
:
theorem
SimpleGraph.Walk.edgeSet_mapLe_eq_edgeSet
{V : Type u}
{G G' : SimpleGraph V}
(h : G ≤ G')
{u v : V}
(p : G.Walk u v)
:
Transferring between graphs #
def
SimpleGraph.Walk.transfer
{V : Type u}
{G : SimpleGraph V}
{u v : V}
(p : G.Walk u v)
(H : SimpleGraph V)
(h : ∀ e ∈ p.edges, e ∈ H.edgeSet)
:
H.Walk u v
The walk p transferred to lie in H, given that H contains its edges.
Equations
- SimpleGraph.Walk.nil.transfer H h_2 = SimpleGraph.Walk.nil
- (SimpleGraph.Walk.cons h_2 p_2).transfer H h_3 = SimpleGraph.Walk.cons ⋯ (p_2.transfer H ⋯)
Instances For
theorem
SimpleGraph.Walk.transfer_self
{V : Type u}
{G : SimpleGraph V}
{u v : V}
(p : G.Walk u v)
:
theorem
SimpleGraph.Walk.transfer_eq_map_ofLE
{V : Type u}
{G : SimpleGraph V}
{u v : V}
(p : G.Walk u v)
{H : SimpleGraph V}
(hp : ∀ e ∈ p.edges, e ∈ H.edgeSet)
(GH : G ≤ H)
:
@[simp]
theorem
SimpleGraph.Walk.edges_transfer
{V : Type u}
{G : SimpleGraph V}
{u v : V}
(p : G.Walk u v)
{H : SimpleGraph V}
(hp : ∀ e ∈ p.edges, e ∈ H.edgeSet)
:
@[simp]
theorem
SimpleGraph.Walk.edgeSet_transfer
{V : Type u}
{G : SimpleGraph V}
{u v : V}
(p : G.Walk u v)
{H : SimpleGraph V}
(hp : ∀ e ∈ p.edges, e ∈ H.edgeSet)
:
@[simp]
theorem
SimpleGraph.Walk.support_transfer
{V : Type u}
{G : SimpleGraph V}
{u v : V}
(p : G.Walk u v)
{H : SimpleGraph V}
(hp : ∀ e ∈ p.edges, e ∈ H.edgeSet)
:
@[simp]
theorem
SimpleGraph.Walk.length_transfer
{V : Type u}
{G : SimpleGraph V}
{u v : V}
(p : G.Walk u v)
{H : SimpleGraph V}
(hp : ∀ e ∈ p.edges, e ∈ H.edgeSet)
:
@[simp]
theorem
SimpleGraph.Walk.transfer_transfer
{V : Type u}
{G : SimpleGraph V}
{u v : V}
(p : G.Walk u v)
{H : SimpleGraph V}
(hp : ∀ e ∈ p.edges, e ∈ H.edgeSet)
{K : SimpleGraph V}
(hp' : ∀ e ∈ (p.transfer H hp).edges, e ∈ K.edgeSet)
:
@[simp]
theorem
SimpleGraph.Walk.reverse_transfer
{V : Type u}
{G : SimpleGraph V}
{u v : V}
(p : G.Walk u v)
{H : SimpleGraph V}
(hp : ∀ e ∈ p.edges, e ∈ H.edgeSet)
:
Inducing a walk #
def
SimpleGraph.Walk.induce
{V : Type u}
{G : SimpleGraph V}
(s : Set V)
{u v : V}
(w : G.Walk u v)
(hw : ∀ x ∈ w.support, x ∈ s)
:
A walk in G which is fully contained in a set s of vertices lifts to a walk of G[s].
Equations
Instances For
@[simp]
theorem
SimpleGraph.Walk.induce_nil
{V : Type u}
{G : SimpleGraph V}
{u : V}
{s : Set V}
(hw : ∀ x ∈ nil.support, x ∈ s)
:
@[simp]
theorem
SimpleGraph.Walk.induce_cons
{V : Type u}
{G : SimpleGraph V}
{u v u' : V}
{s : Set V}
(huu' : G.Adj u u')
(w : G.Walk u' v)
(hw : ∀ x ∈ (cons huu' w).support, x ∈ s)
:
@[simp]
theorem
SimpleGraph.Walk.support_induce
{V : Type u}
{G : SimpleGraph V}
{s : Set V}
{u v : V}
(w : G.Walk u v)
(hw : ∀ x ∈ w.support, x ∈ s)
:
@[simp]
theorem
SimpleGraph.Walk.map_induce
{V : Type u}
{G : SimpleGraph V}
{s : Set V}
{u v : V}
(w : G.Walk u v)
(hw : ∀ x ∈ w.support, x ∈ s)
:
theorem
SimpleGraph.Walk.map_induce_induceHomOfLE
{V : Type u}
{G : SimpleGraph V}
{s s' : Set V}
(hs : s ⊆ s')
{u v : V}
(w : G.Walk u v)
(hw : ∀ x ∈ w.support, x ∈ s)
:
Deleting edges #
@[reducible, inline]
abbrev
SimpleGraph.Walk.toDeleteEdges
{V : Type u}
{G : SimpleGraph V}
(s : Set (Sym2 V))
{v w : V}
(p : G.Walk v w)
(hp : ∀ e ∈ p.edges, e ∉ s)
:
(G.deleteEdges s).Walk v w
Given a walk that avoids a set of edges, produce a walk in the graph with those edges deleted.
Equations
- SimpleGraph.Walk.toDeleteEdges s p hp = p.transfer (G.deleteEdges s) ⋯
Instances For
@[simp]
theorem
SimpleGraph.Walk.toDeleteEdges_nil
{V : Type u}
{G : SimpleGraph V}
(s : Set (Sym2 V))
{v : V}
(hp : ∀ e ∈ nil.edges, e ∉ s)
:
@[simp]
theorem
SimpleGraph.Walk.toDeleteEdges_cons
{V : Type u}
{G : SimpleGraph V}
(s : Set (Sym2 V))
{u v w : V}
(h : G.Adj u v)
(p : G.Walk v w)
(hp : ∀ e ∈ (cons h p).edges, e ∉ s)
:
@[reducible, inline]
abbrev
SimpleGraph.Walk.toDeleteEdge
{V : Type u}
{G : SimpleGraph V}
{v w : V}
(e : Sym2 V)
(p : G.Walk v w)
(hp : e ∉ p.edges)
:
(G.deleteEdges {e}).Walk v w
Given a walk that avoids an edge, create a walk in the subgraph with that edge deleted.
This is an abbreviation for SimpleGraph.Walk.toDeleteEdges.
Equations
Instances For
@[simp]
theorem
SimpleGraph.Walk.map_toDeleteEdges_eq
{V : Type u}
{G : SimpleGraph V}
{v w : V}
(s : Set (Sym2 V))
{p : G.Walk v w}
(hp : ∀ e ∈ p.edges, e ∉ s)
: