Documentation

Mathlib.Combinatorics.SimpleGraph.Subgraph

Subgraphs of a simple graph #

A subgraph of a simple graph consists of subsets of the graph's vertices and edges such that the endpoints of each edge are present in the vertex subset. The edge subset is formalized as a sub-relation of the adjacency relation of the simple graph.

Main definitions #

Implementation notes #

TODO #

structure SimpleGraph.Subgraph {V : Type u} (G : SimpleGraph V) :

A subgraph of a SimpleGraph is a subset of vertices along with a restriction of the adjacency relation that is symmetric and is supported by the vertex subset. They also form a bounded lattice.

Thinking of V → V → Prop as Set (V × V), a set of darts (i.e., half-edges), then Subgraph.adj_sub is that the darts of a subgraph are a subset of the darts of G.

Instances For
    theorem SimpleGraph.Subgraph.ext {V : Type u} {G : SimpleGraph V} {x y : G.Subgraph} (verts : x.verts = y.verts) (Adj : x.Adj = y.Adj) :
    x = y

    The one-vertex subgraph.

    Equations
    Instances For
      @[simp]
      @[simp]
      theorem SimpleGraph.singletonSubgraph_adj {V : Type u} (G : SimpleGraph V) (v a✝ a✝¹ : V) :
      (G.singletonSubgraph v).Adj a✝ a✝¹ = a✝ a✝¹
      def SimpleGraph.subgraphOfAdj {V : Type u} (G : SimpleGraph V) {v w : V} (hvw : G.Adj v w) :

      The one-edge subgraph.

      Equations
      • G.subgraphOfAdj hvw = { verts := {v, w}, Adj := fun (a b : V) => s(v, w) = s(a, b), adj_sub := , edge_vert := , symm := }
      Instances For
        @[simp]
        theorem SimpleGraph.subgraphOfAdj_adj {V : Type u} (G : SimpleGraph V) {v w : V} (hvw : G.Adj v w) (a b : V) :
        (G.subgraphOfAdj hvw).Adj a b = (s(v, w) = s(a, b))
        @[simp]
        theorem SimpleGraph.subgraphOfAdj_verts {V : Type u} (G : SimpleGraph V) {v w : V} (hvw : G.Adj v w) :
        theorem SimpleGraph.Subgraph.adj_comm {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (v w : V) :
        G'.Adj v w G'.Adj w v
        theorem SimpleGraph.Subgraph.adj_symm {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) {u v : V} (h : G'.Adj u v) :
        G'.Adj v u
        theorem SimpleGraph.Subgraph.Adj.symm {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {u v : V} (h : G'.Adj u v) :
        G'.Adj v u
        theorem SimpleGraph.Subgraph.Adj.adj_sub {V : Type u} {G : SimpleGraph V} {H : G.Subgraph} {u v : V} (h : H.Adj u v) :
        G.Adj u v
        theorem SimpleGraph.Subgraph.Adj.fst_mem {V : Type u} {G : SimpleGraph V} {H : G.Subgraph} {u v : V} (h : H.Adj u v) :
        theorem SimpleGraph.Subgraph.Adj.snd_mem {V : Type u} {G : SimpleGraph V} {H : G.Subgraph} {u v : V} (h : H.Adj u v) :
        theorem SimpleGraph.Subgraph.Adj.ne {V : Type u} {G : SimpleGraph V} {H : G.Subgraph} {u v : V} (h : H.Adj u v) :
        u v
        theorem SimpleGraph.Subgraph.adj_congr_of_sym2 {V : Type u} {G : SimpleGraph V} {H : G.Subgraph} {u v w x : V} (h2 : s(u, v) = s(w, x)) :
        H.Adj u v H.Adj w x

        Coercion from G' : Subgraph G to a SimpleGraph G'.verts.

        Equations
        • G'.coe = { Adj := fun (v w : G'.verts) => G'.Adj v w, symm := , loopless := }
        Instances For
          @[simp]
          theorem SimpleGraph.Subgraph.coe_adj {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (v w : G'.verts) :
          G'.coe.Adj v w = G'.Adj v w
          @[simp]
          theorem SimpleGraph.Subgraph.coe_adj_sub {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (u v : G'.verts) (h : G'.coe.Adj u v) :
          G.Adj u v
          theorem SimpleGraph.Subgraph.Adj.coe {V : Type u} {G : SimpleGraph V} {H : G.Subgraph} {u v : V} (h : H.Adj u v) :
          H.coe.Adj u, v,

          A subgraph is called a spanning subgraph if it contains all the vertices of G.

          Equations
          Instances For

            Coercion from Subgraph G to SimpleGraph V. If G' is a spanning subgraph, then G'.spanningCoe yields an isomorphic graph. In general, this adds in all vertices from V as isolated vertices.

            Equations
            Instances For
              @[simp]
              theorem SimpleGraph.Subgraph.spanningCoe_adj {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (a✝ a✝¹ : V) :
              G'.spanningCoe.Adj a✝ a✝¹ = G'.Adj a✝ a✝¹
              @[simp]
              theorem SimpleGraph.Subgraph.Adj.of_spanningCoe {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {u v : G'.verts} (h : G'.spanningCoe.Adj u v) :
              G.Adj u v
              theorem SimpleGraph.Subgraph.spanningCoe_inj {V : Type u} {G : SimpleGraph V} {G₁ G₂ : G.Subgraph} :
              G₁.spanningCoe = G₂.spanningCoe G₁.Adj = G₂.Adj
              theorem SimpleGraph.Subgraph.mem_of_adj_spanningCoe {V : Type u} {v w : V} {s : Set V} (G : SimpleGraph s) (hadj : G.spanningCoe.Adj v w) :
              v s
              @[simp]

              spanningCoe is equivalent to coe for a subgraph that IsSpanning.

              Equations
              Instances For

                A subgraph is called an induced subgraph if vertices of G' are adjacent if they are adjacent in G.

                Equations
                Instances For

                  H.support is the set of vertices that form edges in the subgraph H.

                  Equations
                  Instances For
                    theorem SimpleGraph.Subgraph.mem_support {V : Type u} {G : SimpleGraph V} (H : G.Subgraph) {v : V} :
                    v H.support ∃ (w : V), H.Adj v w
                    def SimpleGraph.Subgraph.neighborSet {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (v : V) :
                    Set V

                    G'.neighborSet v is the set of vertices adjacent to v in G'.

                    Equations
                    Instances For
                      @[simp]
                      theorem SimpleGraph.Subgraph.mem_neighborSet {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (v w : V) :
                      w G'.neighborSet v G'.Adj v w
                      def SimpleGraph.Subgraph.coeNeighborSetEquiv {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} (v : G'.verts) :
                      (G'.coe.neighborSet v) (G'.neighborSet v)

                      A subgraph as a graph has equivalent neighbor sets.

                      Equations
                      Instances For
                        def SimpleGraph.Subgraph.edgeSet {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) :
                        Set (Sym2 V)

                        The edge set of G' consists of a subset of edges of G.

                        Equations
                        Instances For
                          @[simp]
                          theorem SimpleGraph.Subgraph.mem_edgeSet {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {v w : V} :
                          s(v, w) G'.edgeSet G'.Adj v w
                          theorem SimpleGraph.Subgraph.mem_verts_of_mem_edge {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {e : Sym2 V} {v : V} (he : e G'.edgeSet) (hv : v e) :
                          v G'.verts
                          @[deprecated SimpleGraph.Subgraph.mem_verts_of_mem_edge (since := "2024-10-01")]
                          theorem SimpleGraph.Subgraph.mem_verts_if_mem_edge {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {e : Sym2 V} {v : V} (he : e G'.edgeSet) (hv : v e) :
                          v G'.verts

                          Alias of SimpleGraph.Subgraph.mem_verts_of_mem_edge.

                          def SimpleGraph.Subgraph.incidenceSet {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (v : V) :
                          Set (Sym2 V)

                          The incidenceSet is the set of edges incident to a given vertex.

                          Equations
                          Instances For
                            @[reducible, inline]
                            abbrev SimpleGraph.Subgraph.vert {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (v : V) (h : v G'.verts) :
                            G'.verts

                            Give a vertex as an element of the subgraph's vertex type.

                            Equations
                            Instances For
                              def SimpleGraph.Subgraph.copy {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (V'' : Set V) (hV : V'' = G'.verts) (adj' : VVProp) (hadj : adj' = G'.Adj) :

                              Create an equal copy of a subgraph (see copy_eq) with possibly different definitional equalities. See Note [range copy pattern].

                              Equations
                              • G'.copy V'' hV adj' hadj = { verts := V'', Adj := adj', adj_sub := , edge_vert := , symm := }
                              Instances For
                                theorem SimpleGraph.Subgraph.copy_eq {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (V'' : Set V) (hV : V'' = G'.verts) (adj' : VVProp) (hadj : adj' = G'.Adj) :
                                G'.copy V'' hV adj' hadj = G'

                                The union of two subgraphs.

                                Equations

                                The intersection of two subgraphs.

                                Equations

                                The top subgraph is G as a subgraph of itself.

                                Equations

                                The bot subgraph is the subgraph with no vertices or edges.

                                Equations
                                Equations
                                • One or more equations did not get rendered due to their size.
                                Equations
                                • One or more equations did not get rendered due to their size.
                                @[simp]
                                theorem SimpleGraph.Subgraph.sup_adj {V : Type u} {G : SimpleGraph V} {G₁ G₂ : G.Subgraph} {a b : V} :
                                (G₁ G₂).Adj a b G₁.Adj a b G₂.Adj a b
                                @[simp]
                                theorem SimpleGraph.Subgraph.inf_adj {V : Type u} {G : SimpleGraph V} {G₁ G₂ : G.Subgraph} {a b : V} :
                                (G₁ G₂).Adj a b G₁.Adj a b G₂.Adj a b
                                @[simp]
                                theorem SimpleGraph.Subgraph.top_adj {V : Type u} {G : SimpleGraph V} {a b : V} :
                                .Adj a b G.Adj a b
                                @[simp]
                                theorem SimpleGraph.Subgraph.not_bot_adj {V : Type u} {G : SimpleGraph V} {a b : V} :
                                @[simp]
                                theorem SimpleGraph.Subgraph.verts_sup {V : Type u} {G : SimpleGraph V} (G₁ G₂ : G.Subgraph) :
                                (G₁ G₂).verts = G₁.verts G₂.verts
                                @[simp]
                                theorem SimpleGraph.Subgraph.verts_inf {V : Type u} {G : SimpleGraph V} (G₁ G₂ : G.Subgraph) :
                                (G₁ G₂).verts = G₁.verts G₂.verts
                                @[simp]
                                theorem SimpleGraph.Subgraph.sSup_adj {V : Type u} {G : SimpleGraph V} {a b : V} {s : Set G.Subgraph} :
                                (sSup s).Adj a b G_1s, G_1.Adj a b
                                @[simp]
                                theorem SimpleGraph.Subgraph.sInf_adj {V : Type u} {G : SimpleGraph V} {a b : V} {s : Set G.Subgraph} :
                                (sInf s).Adj a b (∀ G's, G'.Adj a b) G.Adj a b
                                @[simp]
                                theorem SimpleGraph.Subgraph.iSup_adj {ι : Sort u_1} {V : Type u} {G : SimpleGraph V} {a b : V} {f : ιG.Subgraph} :
                                (⨆ (i : ι), f i).Adj a b ∃ (i : ι), (f i).Adj a b
                                @[simp]
                                theorem SimpleGraph.Subgraph.iInf_adj {ι : Sort u_1} {V : Type u} {G : SimpleGraph V} {a b : V} {f : ιG.Subgraph} :
                                (⨅ (i : ι), f i).Adj a b (∀ (i : ι), (f i).Adj a b) G.Adj a b
                                theorem SimpleGraph.Subgraph.sInf_adj_of_nonempty {V : Type u} {G : SimpleGraph V} {a b : V} {s : Set G.Subgraph} (hs : s.Nonempty) :
                                (sInf s).Adj a b G's, G'.Adj a b
                                theorem SimpleGraph.Subgraph.iInf_adj_of_nonempty {ι : Sort u_1} {V : Type u} {G : SimpleGraph V} {a b : V} [Nonempty ι] {f : ιG.Subgraph} :
                                (⨅ (i : ι), f i).Adj a b ∀ (i : ι), (f i).Adj a b
                                @[simp]
                                theorem SimpleGraph.Subgraph.verts_sSup {V : Type u} {G : SimpleGraph V} (s : Set G.Subgraph) :
                                (sSup s).verts = G's, G'.verts
                                @[simp]
                                theorem SimpleGraph.Subgraph.verts_sInf {V : Type u} {G : SimpleGraph V} (s : Set G.Subgraph) :
                                (sInf s).verts = G's, G'.verts
                                @[simp]
                                theorem SimpleGraph.Subgraph.verts_iSup {ι : Sort u_1} {V : Type u} {G : SimpleGraph V} {f : ιG.Subgraph} :
                                (⨆ (i : ι), f i).verts = ⋃ (i : ι), (f i).verts
                                @[simp]
                                theorem SimpleGraph.Subgraph.verts_iInf {ι : Sort u_1} {V : Type u} {G : SimpleGraph V} {f : ιG.Subgraph} :
                                (⨅ (i : ι), f i).verts = ⋂ (i : ι), (f i).verts

                                The graph isomorphism between the top element of G.subgraph and G.

                                Equations
                                Instances For

                                  For subgraphs G₁, G₂, G₁ ≤ G₂ iff G₁.verts ⊆ G₂.verts and ∀ a b, G₁.adj a b → G₂.adj a b.

                                  Equations

                                  Note that subgraphs do not form a Boolean algebra, because of verts.

                                  Equations
                                  Instances For
                                    theorem SimpleGraph.Subgraph.verts_mono {V : Type u} {G : SimpleGraph V} {H H' : G.Subgraph} (h : H H') :
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.neighborSet_sup {V : Type u} {G : SimpleGraph V} {H H' : G.Subgraph} (v : V) :
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.neighborSet_inf {V : Type u} {G : SimpleGraph V} {H H' : G.Subgraph} (v : V) :
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.neighborSet_sSup {V : Type u} {G : SimpleGraph V} (s : Set G.Subgraph) (v : V) :
                                    (sSup s).neighborSet v = G's, G'.neighborSet v
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.neighborSet_sInf {V : Type u} {G : SimpleGraph V} (s : Set G.Subgraph) (v : V) :
                                    (sInf s).neighborSet v = (⋂ G's, G'.neighborSet v) G.neighborSet v
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.neighborSet_iSup {ι : Sort u_1} {V : Type u} {G : SimpleGraph V} (f : ιG.Subgraph) (v : V) :
                                    (⨆ (i : ι), f i).neighborSet v = ⋃ (i : ι), (f i).neighborSet v
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.neighborSet_iInf {ι : Sort u_1} {V : Type u} {G : SimpleGraph V} (f : ιG.Subgraph) (v : V) :
                                    (⨅ (i : ι), f i).neighborSet v = (⋂ (i : ι), (f i).neighborSet v) G.neighborSet v
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.edgeSet_inf {V : Type u} {G : SimpleGraph V} {H₁ H₂ : G.Subgraph} :
                                    (H₁ H₂).edgeSet = H₁.edgeSet H₂.edgeSet
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.edgeSet_sup {V : Type u} {G : SimpleGraph V} {H₁ H₂ : G.Subgraph} :
                                    (H₁ H₂).edgeSet = H₁.edgeSet H₂.edgeSet
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.edgeSet_sSup {V : Type u} {G : SimpleGraph V} (s : Set G.Subgraph) :
                                    (sSup s).edgeSet = G's, G'.edgeSet
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.edgeSet_sInf {V : Type u} {G : SimpleGraph V} (s : Set G.Subgraph) :
                                    (sInf s).edgeSet = (⋂ G's, G'.edgeSet) G.edgeSet
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.edgeSet_iSup {ι : Sort u_1} {V : Type u} {G : SimpleGraph V} (f : ιG.Subgraph) :
                                    (⨆ (i : ι), f i).edgeSet = ⋃ (i : ι), (f i).edgeSet
                                    @[simp]
                                    theorem SimpleGraph.Subgraph.edgeSet_iInf {ι : Sort u_1} {V : Type u} {G : SimpleGraph V} (f : ιG.Subgraph) :
                                    (⨅ (i : ι), f i).edgeSet = (⋂ (i : ι), (f i).edgeSet) G.edgeSet
                                    def SimpleGraph.toSubgraph {V : Type u} {G : SimpleGraph V} (H : SimpleGraph V) (h : H G) :

                                    Turn a subgraph of a SimpleGraph into a member of its subgraph type.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem SimpleGraph.toSubgraph_adj {V : Type u} {G : SimpleGraph V} (H : SimpleGraph V) (h : H G) (a✝ a✝¹ : V) :
                                      (toSubgraph H h).Adj a✝ a✝¹ = H.Adj a✝ a✝¹
                                      @[simp]
                                      theorem SimpleGraph.toSubgraph_verts {V : Type u} {G : SimpleGraph V} (H : SimpleGraph V) (h : H G) :
                                      theorem SimpleGraph.Subgraph.support_mono {V : Type u} {G : SimpleGraph V} {H H' : G.Subgraph} (h : H H') :

                                      The top of the Subgraph G lattice is equivalent to the graph itself.

                                      Equations
                                      Instances For

                                        The bottom of the Subgraph G lattice is equivalent to the empty graph on the empty vertex type.

                                        Equations
                                        Instances For
                                          theorem SimpleGraph.Subgraph.edgeSet_mono {V : Type u} {G : SimpleGraph V} {H₁ H₂ : G.Subgraph} (h : H₁ H₂) :
                                          H₁.edgeSet H₂.edgeSet
                                          theorem Disjoint.edgeSet {V : Type u} {G : SimpleGraph V} {H₁ H₂ : G.Subgraph} (h : Disjoint H₁ H₂) :
                                          def SimpleGraph.Subgraph.map {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') (H : G.Subgraph) :

                                          Graph homomorphisms induce a covariant function on subgraphs.

                                          Equations
                                          Instances For
                                            @[simp]
                                            theorem SimpleGraph.Subgraph.map_adj {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') (H : G.Subgraph) (a✝ a✝¹ : W) :
                                            (Subgraph.map f H).Adj a✝ a✝¹ = Relation.Map H.Adj (⇑f) (⇑f) a✝ a✝¹
                                            @[simp]
                                            theorem SimpleGraph.Subgraph.map_verts {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') (H : G.Subgraph) :
                                            (Subgraph.map f H).verts = f '' H.verts
                                            @[simp]
                                            theorem SimpleGraph.Subgraph.map_comp {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} {U : Type u_2} {G'' : SimpleGraph U} (H : G.Subgraph) (f : G →g G') (g : G' →g G'') :
                                            theorem SimpleGraph.Subgraph.map_mono {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} {f : G →g G'} {H₁ H₂ : G.Subgraph} (hH : H₁ H₂) :
                                            theorem SimpleGraph.Subgraph.map_monotone {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} {f : G →g G'} :
                                            theorem SimpleGraph.Subgraph.map_sup {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') (H₁ H₂ : G.Subgraph) :
                                            Subgraph.map f (H₁ H₂) = Subgraph.map f H₁ Subgraph.map f H₂
                                            def SimpleGraph.Subgraph.comap {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') (H : G'.Subgraph) :

                                            Graph homomorphisms induce a contravariant function on subgraphs.

                                            Equations
                                            Instances For
                                              @[simp]
                                              theorem SimpleGraph.Subgraph.comap_verts {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') (H : G'.Subgraph) :
                                              @[simp]
                                              theorem SimpleGraph.Subgraph.comap_adj {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') (H : G'.Subgraph) (u v : V) :
                                              (Subgraph.comap f H).Adj u v = (G.Adj u v H.Adj (f u) (f v))
                                              theorem SimpleGraph.Subgraph.map_le_iff_le_comap {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') (H : G.Subgraph) (H' : G'.Subgraph) :
                                              def SimpleGraph.Subgraph.inclusion {V : Type u} {G : SimpleGraph V} {x y : G.Subgraph} (h : x y) :

                                              Given two subgraphs, one a subgraph of the other, there is an induced injective homomorphism of the subgraphs as graphs.

                                              Equations
                                              Instances For
                                                @[simp]
                                                theorem SimpleGraph.Subgraph.inclusion_apply_coe {V : Type u} {G : SimpleGraph V} {x y : G.Subgraph} (h : x y) (v : x.verts) :
                                                ((inclusion h) v) = v
                                                def SimpleGraph.Subgraph.hom {V : Type u} {G : SimpleGraph V} (x : G.Subgraph) :
                                                x.coe →g G

                                                There is an induced injective homomorphism of a subgraph of G into G.

                                                Equations
                                                • x.hom = { toFun := fun (v : x.verts) => v, map_rel' := }
                                                Instances For
                                                  @[simp]
                                                  theorem SimpleGraph.Subgraph.hom_apply {V : Type u} {G : SimpleGraph V} (x : G.Subgraph) (v : x.verts) :
                                                  x.hom v = v
                                                  @[simp]
                                                  theorem SimpleGraph.Subgraph.coe_hom {V : Type u} {G : SimpleGraph V} (x : G.Subgraph) :
                                                  x.hom = fun (v : x.verts) => v

                                                  There is an induced injective homomorphism of a subgraph of G as a spanning subgraph into G.

                                                  Equations
                                                  Instances For
                                                    @[simp]
                                                    theorem SimpleGraph.Subgraph.spanningHom_apply {V : Type u} {G : SimpleGraph V} (x : G.Subgraph) (a : V) :
                                                    instance SimpleGraph.Subgraph.finiteAt {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} (v : G'.verts) [DecidableRel G'.Adj] [Fintype (G.neighborSet v)] :
                                                    Fintype (G'.neighborSet v)

                                                    If a graph is locally finite at a vertex, then so is a subgraph of that graph.

                                                    Equations
                                                    def SimpleGraph.Subgraph.finiteAtOfSubgraph {V : Type u} {G : SimpleGraph V} {G' G'' : G.Subgraph} [DecidableRel G'.Adj] (h : G' G'') (v : G'.verts) [Fintype (G''.neighborSet v)] :
                                                    Fintype (G'.neighborSet v)

                                                    If a subgraph is locally finite at a vertex, then so are subgraphs of that subgraph.

                                                    This is not an instance because G'' cannot be inferred.

                                                    Equations
                                                    Instances For
                                                      def SimpleGraph.Subgraph.degree {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (v : V) [Fintype (G'.neighborSet v)] :

                                                      The degree of a vertex in a subgraph. It's zero for vertices outside the subgraph.

                                                      Equations
                                                      Instances For
                                                        theorem SimpleGraph.Subgraph.degree_le {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (v : V) [Fintype (G'.neighborSet v)] [Fintype (G.neighborSet v)] :
                                                        G'.degree v G.degree v
                                                        theorem SimpleGraph.Subgraph.degree_le' {V : Type u} {G : SimpleGraph V} (G' G'' : G.Subgraph) (h : G' G'') (v : V) [Fintype (G'.neighborSet v)] [Fintype (G''.neighborSet v)] :
                                                        G'.degree v G''.degree v
                                                        @[simp]
                                                        theorem SimpleGraph.Subgraph.coe_degree {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (v : G'.verts) [Fintype (G'.coe.neighborSet v)] [Fintype (G'.neighborSet v)] :
                                                        G'.coe.degree v = G'.degree v
                                                        @[simp]
                                                        theorem SimpleGraph.Subgraph.degree_eq_one_iff_unique_adj {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {v : V} [Fintype (G'.neighborSet v)] :
                                                        G'.degree v = 1 ∃! w : V, G'.Adj v w
                                                        theorem SimpleGraph.Subgraph.neighborSet_eq_of_equiv {V : Type u} {G : SimpleGraph V} {v : V} {H : G.Subgraph} (h : (G.neighborSet v) (H.neighborSet v)) (hfin : (G.neighborSet v).Finite) :
                                                        theorem SimpleGraph.Subgraph.adj_iff_of_neighborSet_equiv {V : Type u} {G : SimpleGraph V} {v : V} {H : G.Subgraph} (h : (G.neighborSet v) (H.neighborSet v)) (hfin : (G.neighborSet v).Finite) {w : V} :
                                                        H.Adj v w G.Adj v w

                                                        Properties of singletonSubgraph and subgraphOfAdj #

                                                        @[simp]
                                                        @[simp]
                                                        theorem SimpleGraph.map_singletonSubgraph {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') {v : V} :
                                                        instance SimpleGraph.nonempty_subgraphOfAdj_verts {V : Type u} {G : SimpleGraph V} {v w : V} (hvw : G.Adj v w) :
                                                        @[simp]
                                                        theorem SimpleGraph.edgeSet_subgraphOfAdj {V : Type u} {G : SimpleGraph V} {v w : V} (hvw : G.Adj v w) :
                                                        theorem SimpleGraph.subgraphOfAdj_le_of_adj {V : Type u} {G : SimpleGraph V} {v w : V} (H : G.Subgraph) (h : H.Adj v w) :
                                                        theorem SimpleGraph.subgraphOfAdj_symm {V : Type u} {G : SimpleGraph V} {v w : V} (hvw : G.Adj v w) :
                                                        @[simp]
                                                        theorem SimpleGraph.map_subgraphOfAdj {V : Type u} {W : Type v} {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') {v w : V} (hvw : G.Adj v w) :
                                                        theorem SimpleGraph.neighborSet_subgraphOfAdj_subset {V : Type u} {G : SimpleGraph V} {u v w : V} (hvw : G.Adj v w) :
                                                        @[simp]
                                                        theorem SimpleGraph.neighborSet_fst_subgraphOfAdj {V : Type u} {G : SimpleGraph V} {v w : V} (hvw : G.Adj v w) :
                                                        @[simp]
                                                        theorem SimpleGraph.neighborSet_snd_subgraphOfAdj {V : Type u} {G : SimpleGraph V} {v w : V} (hvw : G.Adj v w) :
                                                        @[simp]
                                                        theorem SimpleGraph.neighborSet_subgraphOfAdj_of_ne_of_ne {V : Type u} {G : SimpleGraph V} {u v w : V} (hvw : G.Adj v w) (hv : u v) (hw : u w) :
                                                        theorem SimpleGraph.neighborSet_subgraphOfAdj {V : Type u} {G : SimpleGraph V} [DecidableEq V] {u v w : V} (hvw : G.Adj v w) :
                                                        @[simp]
                                                        theorem SimpleGraph.support_subgraphOfAdj {V : Type u} {G : SimpleGraph V} {u v : V} (h : G.Adj u v) :

                                                        Subgraphs of subgraphs #

                                                        @[reducible, inline]

                                                        Given a subgraph of a subgraph of G, construct a subgraph of G.

                                                        Equations
                                                        Instances For
                                                          @[reducible, inline]
                                                          abbrev SimpleGraph.Subgraph.restrict {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} :

                                                          Given a subgraph of G, restrict it to being a subgraph of another subgraph G' by taking the portion of G that intersects G'.

                                                          Equations
                                                          Instances For
                                                            theorem SimpleGraph.Subgraph.coeSubgraph_adj {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} (G'' : G'.coe.Subgraph) (v w : V) :
                                                            (Subgraph.coeSubgraph G'').Adj v w ∃ (hv : v G'.verts) (hw : w G'.verts), G''.Adj v, hv w, hw
                                                            theorem SimpleGraph.Subgraph.restrict_adj {V : Type u} {G : SimpleGraph V} {G' G'' : G.Subgraph} (v w : G'.verts) :
                                                            (Subgraph.restrict G'').Adj v w G'.Adj v w G''.Adj v w

                                                            Edge deletion #

                                                            def SimpleGraph.Subgraph.deleteEdges {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (s : Set (Sym2 V)) :

                                                            Given a subgraph G' and a set of vertex pairs, remove all of the corresponding edges from its edge set, if present.

                                                            See also: SimpleGraph.deleteEdges.

                                                            Equations
                                                            Instances For
                                                              @[simp]
                                                              theorem SimpleGraph.Subgraph.deleteEdges_verts {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} (s : Set (Sym2 V)) :
                                                              @[simp]
                                                              theorem SimpleGraph.Subgraph.deleteEdges_adj {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} (s : Set (Sym2 V)) (v w : V) :
                                                              (G'.deleteEdges s).Adj v w G'.Adj v w s(v, w)s
                                                              @[simp]
                                                              theorem SimpleGraph.Subgraph.deleteEdges_deleteEdges {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} (s s' : Set (Sym2 V)) :
                                                              theorem SimpleGraph.Subgraph.deleteEdges_le {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} (s : Set (Sym2 V)) :
                                                              theorem SimpleGraph.Subgraph.deleteEdges_le_of_le {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {s s' : Set (Sym2 V)} (h : s s') :

                                                              Induced subgraphs #

                                                              def SimpleGraph.Subgraph.induce {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (s : Set V) :

                                                              The induced subgraph of a subgraph. The expectation is that s ⊆ G'.verts for the usual notion of an induced subgraph, but, in general, s is taken to be the new vertex set and edges are induced from the subgraph G'.

                                                              Equations
                                                              • G'.induce s = { verts := s, Adj := fun (u v : V) => u s v s G'.Adj u v, adj_sub := , edge_vert := , symm := }
                                                              Instances For
                                                                @[simp]
                                                                theorem SimpleGraph.Subgraph.induce_adj {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (s : Set V) (u v : V) :
                                                                (G'.induce s).Adj u v = (u s v s G'.Adj u v)
                                                                @[simp]
                                                                theorem SimpleGraph.Subgraph.induce_verts {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (s : Set V) :
                                                                (G'.induce s).verts = s
                                                                theorem SimpleGraph.Subgraph.induce_mono {V : Type u} {G : SimpleGraph V} {G' G'' : G.Subgraph} {s s' : Set V} (hg : G' G'') (hs : s s') :
                                                                G'.induce s G''.induce s'
                                                                theorem SimpleGraph.Subgraph.induce_mono_left {V : Type u} {G : SimpleGraph V} {G' G'' : G.Subgraph} {s : Set V} (hg : G' G'') :
                                                                G'.induce s G''.induce s
                                                                theorem SimpleGraph.Subgraph.induce_mono_right {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {s s' : Set V} (hs : s s') :
                                                                G'.induce s G'.induce s'
                                                                @[simp]
                                                                @[simp]
                                                                theorem SimpleGraph.Subgraph.induce_self_verts {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} :
                                                                G'.induce G'.verts = G'
                                                                theorem SimpleGraph.Subgraph.le_induce_union {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {s s' : Set V} :
                                                                G'.induce s G'.induce s' G'.induce (s s')
                                                                theorem SimpleGraph.Subgraph.le_induce_union_left {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {s s' : Set V} :
                                                                G'.induce s G'.induce (s s')
                                                                theorem SimpleGraph.Subgraph.le_induce_union_right {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {s s' : Set V} :
                                                                G'.induce s' G'.induce (s s')
                                                                theorem SimpleGraph.Subgraph.subgraphOfAdj_eq_induce {V : Type u} {G : SimpleGraph V} {v w : V} (hvw : G.Adj v w) :
                                                                @[reducible, inline]
                                                                abbrev SimpleGraph.Subgraph.deleteVerts {V : Type u} {G : SimpleGraph V} (G' : G.Subgraph) (s : Set V) :

                                                                Given a subgraph and a set of vertices, delete all the vertices from the subgraph, if present. Any edges incident to the deleted vertices are deleted as well.

                                                                Equations
                                                                Instances For
                                                                  theorem SimpleGraph.Subgraph.deleteVerts_verts {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {s : Set V} :
                                                                  (G'.deleteVerts s).verts = G'.verts \ s
                                                                  theorem SimpleGraph.Subgraph.deleteVerts_adj {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {s : Set V} {u v : V} :
                                                                  (G'.deleteVerts s).Adj u v u G'.verts us v G'.verts vs G'.Adj u v
                                                                  @[simp]
                                                                  theorem SimpleGraph.Subgraph.deleteVerts_deleteVerts {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} (s s' : Set V) :
                                                                  theorem SimpleGraph.Subgraph.deleteVerts_le {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {s : Set V} :
                                                                  theorem SimpleGraph.Subgraph.deleteVerts_mono {V : Type u} {G : SimpleGraph V} {s : Set V} {G' G'' : G.Subgraph} (h : G' G'') :
                                                                  theorem SimpleGraph.Subgraph.deleteVerts_anti {V : Type u} {G : SimpleGraph V} {G' : G.Subgraph} {s s' : Set V} (h : s s') :