Skip to content

Commit

Permalink
feat(combinatorics/simple_graph/connectivity): define connected compo…
Browse files Browse the repository at this point in the history
…nents (#12766)
  • Loading branch information
kmill committed Apr 14, 2022
1 parent 251bd84 commit cbf3062
Showing 1 changed file with 64 additions and 2 deletions.
66 changes: 64 additions & 2 deletions src/combinatorics/simple_graph/connectivity.lean
Expand Up @@ -52,6 +52,9 @@ counterparts in [Chou1994].
* `simple_graph.subgraph.connected` gives subgraphs the connectivity
predicate via `simple_graph.subgraph.coe`.
* `simple_graph.connected_component` is the type of connected components of
a given graph.
## Tags
walks, trails, paths, circuits, cycles
Expand Down Expand Up @@ -988,11 +991,70 @@ structure connected : Prop :=
instance : has_coe_to_fun G.connected (λ _, Π (u v : V), G.reachable u v) :=
⟨λ h, h.preconnected⟩

/-- A subgraph is connected if it is connected as a simple graph. -/
abbreviation subgraph.connected {G : simple_graph V} (H : G.subgraph) : Prop := H.coe.connected
/-- The quotient of `V` by the `simple_graph.reachable` relation gives the connected
components of a graph. -/
def connected_component := quot G.reachable

/-- Gives the connected component containing a particular vertex. -/
def connected_component_mk (v : V) : G.connected_component := quot.mk G.reachable v

instance connected_component.inhabited [inhabited V] : inhabited G.connected_component :=
⟨G.connected_component_mk default⟩

section connected_component
variables {G}

@[elab_as_eliminator]
protected lemma connected_component.ind {β : G.connected_component → Prop}
(h : ∀ (v : V), β (G.connected_component_mk v)) (c : G.connected_component) : β c :=
quot.ind h c

@[elab_as_eliminator]
protected lemma connected_component.ind₂ {β : G.connected_component → G.connected_component → Prop}
(h : ∀ (v w : V), β (G.connected_component_mk v) (G.connected_component_mk w))
(c d : G.connected_component) : β c d :=
quot.induction_on₂ c d h

protected lemma connected_component.sound {v w : V} :
G.reachable v w → G.connected_component_mk v = G.connected_component_mk w := quot.sound

protected lemma connected_component.exact {v w : V} :
G.connected_component_mk v = G.connected_component_mk w → G.reachable v w :=
@quotient.exact _ G.reachable_setoid _ _

@[simp] protected lemma connected_component.eq {v w : V} :
G.connected_component_mk v = G.connected_component_mk w ↔ G.reachable v w :=
@quotient.eq _ G.reachable_setoid _ _

/-- The `connected_component` specialization of `quot.lift`. Provides the stronger
assumption that the vertices are connected by a path. -/
protected def connected_component.lift {β : Sort*} (f : V → β)
(h : ∀ (v w : V) (p : G.walk v w), p.is_path → f v = f w) : G.connected_component → β :=
quot.lift f (λ v w (h' : G.reachable v w), h'.elim_path (λ hp, h v w hp hp.2))

@[simp] protected lemma connected_component.lift_mk {β : Sort*} {f : V → β}
{h : ∀ (v w : V) (p : G.walk v w), p.is_path → f v = f w} {v : V} :
connected_component.lift f h (G.connected_component_mk v) = f v := rfl

protected lemma connected_componentexists» {p : G.connected_component → Prop} :
(∃ (c : G.connected_component), p c) ↔ ∃ v, p (G.connected_component_mk v) :=
(surjective_quot_mk G.reachable).exists

protected lemma connected_componentforall» {p : G.connected_component → Prop} :
(∀ (c : G.connected_component), p c) ↔ ∀ v, p (G.connected_component_mk v) :=
(surjective_quot_mk G.reachable).forall

lemma preconnected.subsingleton_connected_component (h : G.preconnected) :
subsingleton G.connected_component :=
⟨connected_component.ind₂ (λ v w, connected_component.sound (h v w))⟩

end connected_component

variables {G}

/-- A subgraph is connected if it is connected as a simple graph. -/
abbreviation subgraph.connected (H : G.subgraph) : Prop := H.coe.connected

lemma preconnected.set_univ_walk_nonempty (hconn : G.preconnected) (u v : V) :
(set.univ : set (G.walk u v)).nonempty :=
by { rw ← set.nonempty_iff_univ_nonempty, exact hconn u v }
Expand Down

0 comments on commit cbf3062

Please sign in to comment.