Documentation

Mathlib.Combinatorics.SimpleGraph.DegreeSum

Degree-sum formula and handshaking lemma #

The degree-sum formula is that the sum of the degrees of the vertices in a finite graph is equal to twice the number of edges. The handshaking lemma, a corollary, is that the number of odd-degree vertices is even.

Main definitions #

Implementation notes #

We give a combinatorial proof by using the facts that (1) the map from darts to vertices is such that each fiber has cardinality the degree of the corresponding vertex and that (2) the map from darts to edges is 2-to-1.

Tags #

simple graphs, sums, degree-sum formula, handshaking lemma

theorem SimpleGraph.dart_fst_fiber {V : Type u} (G : SimpleGraph V) [Fintype V] [DecidableRel G.Adj] [DecidableEq V] (v : V) :
Finset.filter (fun (d : SimpleGraph.Dart G) => d.toProd.1 = v) Finset.univ = Finset.image (SimpleGraph.dartOfNeighborSet G v) Finset.univ
theorem SimpleGraph.dart_fst_fiber_card_eq_degree {V : Type u} (G : SimpleGraph V) [Fintype V] [DecidableRel G.Adj] [DecidableEq V] (v : V) :
(Finset.filter (fun (d : SimpleGraph.Dart G) => d.toProd.1 = v) Finset.univ).card = SimpleGraph.degree G v
theorem SimpleGraph.dart_edge_fiber_card {V : Type u} (G : SimpleGraph V) [Fintype V] [DecidableRel G.Adj] [DecidableEq V] (e : Sym2 V) (h : e SimpleGraph.edgeSet G) :
(Finset.filter (fun (d : SimpleGraph.Dart G) => SimpleGraph.Dart.edge d = e) Finset.univ).card = 2
theorem SimpleGraph.sum_degrees_eq_twice_card_edges {V : Type u} (G : SimpleGraph V) [Fintype V] [DecidableRel G.Adj] [Fintype (Sym2 V)] :
(Finset.sum Finset.univ fun (v : V) => SimpleGraph.degree G v) = 2 * (SimpleGraph.edgeFinset G).card

The degree-sum formula. This is also known as the handshaking lemma, which might more specifically refer to SimpleGraph.even_card_odd_degree_vertices.

theorem SimpleGraph.even_card_odd_degree_vertices {V : Type u} (G : SimpleGraph V) [Fintype V] [DecidableRel G.Adj] :
Even (Finset.filter (fun (v : V) => Odd (SimpleGraph.degree G v)) Finset.univ).card

The handshaking lemma. See also SimpleGraph.sum_degrees_eq_twice_card_edges.

theorem SimpleGraph.odd_card_odd_degree_vertices_ne {V : Type u} (G : SimpleGraph V) [Fintype V] [DecidableEq V] [DecidableRel G.Adj] (v : V) (h : Odd (SimpleGraph.degree G v)) :
Odd (Finset.filter (fun (w : V) => w v Odd (SimpleGraph.degree G w)) Finset.univ).card