Documentation

Mathlib.Combinatorics.SimpleGraph.Walks.Maps

Mapping walks between graphs #

Functions that map walks between different graphs.

Main definitions #

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} :
G.Walk u vG'.Walk (f u) (f 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) :
    Walk.map f (cons h p) = cons (Walk.map f p)
    @[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') :
    Walk.map f (p.copy hu hv) = (Walk.map f p).copy
    @[simp]
    theorem SimpleGraph.Walk.map_id {V : Type u} {G : SimpleGraph V} {u v : V} (p : G.Walk u v) :
    @[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) :
    Walk.map f' (Walk.map f p) = Walk.map (f'.comp f) p
    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') :
    Walk.map f p = (Walk.map f' p).copy

    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) :
    Walk.map f (p.append q) = (Walk.map f p).append (Walk.map f q)
    @[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 : ) :
    (Walk.map f p).getVert n = f (p.getVert 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) :
      (mapLe h p).edges = p.edges
      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 : ep.edges, e H.edgeSet) :
      H.Walk u v

      The walk p transferred to lie in H, given that H contains its edges.

      Equations
      Instances For
        theorem SimpleGraph.Walk.transfer_self {V : Type u} {G : SimpleGraph V} {u v : V} (p : G.Walk u v) :
        p.transfer G = p
        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 : ep.edges, e H.edgeSet) (GH : G H) :
        p.transfer H hp = Walk.map (Hom.ofLE GH) p
        @[simp]
        theorem SimpleGraph.Walk.edges_transfer {V : Type u} {G : SimpleGraph V} {u v : V} (p : G.Walk u v) {H : SimpleGraph V} (hp : ep.edges, e H.edgeSet) :
        (p.transfer H hp).edges = p.edges
        @[simp]
        theorem SimpleGraph.Walk.edgeSet_transfer {V : Type u} {G : SimpleGraph V} {u v : V} (p : G.Walk u v) {H : SimpleGraph V} (hp : ep.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 : ep.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 : ep.edges, e H.edgeSet) :
        (p.transfer H hp).length = p.length
        @[simp]
        theorem SimpleGraph.Walk.transfer_transfer {V : Type u} {G : SimpleGraph V} {u v : V} (p : G.Walk u v) {H : SimpleGraph V} (hp : ep.edges, e H.edgeSet) {K : SimpleGraph V} (hp' : e(p.transfer H hp).edges, e K.edgeSet) :
        (p.transfer H hp).transfer K hp' = p.transfer K
        @[simp]
        theorem SimpleGraph.Walk.transfer_append {V : Type u} {G : SimpleGraph V} {u v : V} (p : G.Walk u v) {H : SimpleGraph V} {w : V} (q : G.Walk v w) (hpq : e(p.append q).edges, e H.edgeSet) :
        (p.append q).transfer H hpq = (p.transfer H ).append (q.transfer H )
        @[simp]
        theorem SimpleGraph.Walk.reverse_transfer {V : Type u} {G : SimpleGraph V} {u v : V} (p : G.Walk u v) {H : SimpleGraph V} (hp : ep.edges, e H.edgeSet) :
        (p.transfer H hp).reverse = p.reverse.transfer H

        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 : xw.support, x s) :
        (induce s G).Walk u, v,

        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 : xnil.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) :
          Walk.induce s (cons huu' w) hw = cons (Walk.induce s w )
          @[simp]
          theorem SimpleGraph.Walk.support_induce {V : Type u} {G : SimpleGraph V} {s : Set V} {u v : V} (w : G.Walk u v) (hw : xw.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 : xw.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 : xw.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 : ep.edges, es) :
          (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
          Instances For
            @[simp]
            theorem SimpleGraph.Walk.toDeleteEdges_nil {V : Type u} {G : SimpleGraph V} (s : Set (Sym2 V)) {v : V} (hp : enil.edges, es) :
            @[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, es) :
            toDeleteEdges s (cons h p) hp = cons (toDeleteEdges s p )
            @[reducible, inline]
            abbrev SimpleGraph.Walk.toDeleteEdge {V : Type u} {G : SimpleGraph V} {v w : V} (e : Sym2 V) (p : G.Walk v w) (hp : ep.edges) :

            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 : ep.edges, es) :
              Walk.map (Hom.ofLE ) (toDeleteEdges s p hp) = p