Documentation

Mathlib.Combinatorics.SimpleGraph.Clique

Graph cliques #

This file defines cliques in simple graphs. A clique is a set of vertices that are pairwise adjacent.

Main declarations #

Cliques #

@[reducible, inline]
abbrev SimpleGraph.IsClique {α : Type u_1} (G : SimpleGraph α) (s : Set α) :

A clique in a graph is a set of vertices that are pairwise adjacent.

Equations
    Instances For
      theorem SimpleGraph.isClique_iff {α : Type u_1} (G : SimpleGraph α) {s : Set α} :
      theorem SimpleGraph.not_isClique_iff {α : Type u_1} (G : SimpleGraph α) {s : Set α} :
      ¬G.IsClique s ∃ (v : s) (w : s), v w ¬G.Adj v w
      theorem SimpleGraph.isClique_iff_induce_eq {α : Type u_1} (G : SimpleGraph α) {s : Set α} :

      A clique is a set of vertices whose induced graph is complete.

      theorem SimpleGraph.isClique_singleton {α : Type u_1} {G : SimpleGraph α} (a : α) :
      theorem SimpleGraph.IsClique.of_subsingleton {α : Type u_1} {s : Set α} {G : SimpleGraph α} (hs : s.Subsingleton) :
      theorem SimpleGraph.isClique_pair {α : Type u_1} {G : SimpleGraph α} {a b : α} :
      G.IsClique {a, b} a bG.Adj a b
      @[simp]
      theorem SimpleGraph.isClique_insert {α : Type u_1} {G : SimpleGraph α} {s : Set α} {a : α} :
      G.IsClique (insert a s) G.IsClique s bs, a bG.Adj a b
      theorem SimpleGraph.isClique_insert_of_notMem {α : Type u_1} {G : SimpleGraph α} {s : Set α} {a : α} (ha : as) :
      G.IsClique (insert a s) G.IsClique s bs, G.Adj a b
      @[deprecated SimpleGraph.isClique_insert_of_notMem (since := "2025-05-23")]
      theorem SimpleGraph.isClique_insert_of_not_mem {α : Type u_1} {G : SimpleGraph α} {s : Set α} {a : α} (ha : as) :
      G.IsClique (insert a s) G.IsClique s bs, G.Adj a b

      Alias of SimpleGraph.isClique_insert_of_notMem.

      theorem SimpleGraph.IsClique.insert {α : Type u_1} {G : SimpleGraph α} {s : Set α} {a : α} (hs : G.IsClique s) (h : bs, a bG.Adj a b) :
      theorem SimpleGraph.IsClique.mono {α : Type u_1} {G H : SimpleGraph α} {s : Set α} (h : G H) :
      G.IsClique sH.IsClique s
      theorem SimpleGraph.IsClique.subset {α : Type u_1} {G : SimpleGraph α} {s t : Set α} (h : t s) :
      G.IsClique sG.IsClique t
      @[simp]
      theorem SimpleGraph.IsClique.subsingleton {α : Type u_1} {s : Set α} :

      Alias of the forward direction of SimpleGraph.isClique_bot_iff.

      theorem SimpleGraph.IsClique.map {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {s : Set α} (h : G.IsClique s) {f : α β} :
      (SimpleGraph.map f G).IsClique (f '' s)
      theorem SimpleGraph.isClique_map_iff_of_nontrivial {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {f : α β} {t : Set β} (ht : t.Nontrivial) :
      (SimpleGraph.map f G).IsClique t ∃ (s : Set α), G.IsClique s f '' s = t
      theorem SimpleGraph.isClique_map_iff {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {f : α β} {t : Set β} :
      (SimpleGraph.map f G).IsClique t t.Subsingleton ∃ (s : Set α), G.IsClique s f '' s = t
      @[simp]
      theorem SimpleGraph.isClique_map_image_iff {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {s : Set α} {f : α β} :
      theorem SimpleGraph.isClique_map_finset_iff_of_nontrivial {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {f : α β} {t : Finset β} (ht : t.Nontrivial) :
      (SimpleGraph.map f G).IsClique t ∃ (s : Finset α), G.IsClique s Finset.map f s = t
      theorem SimpleGraph.isClique_map_finset_iff {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {f : α β} {t : Finset β} :
      (SimpleGraph.map f G).IsClique t t.card 1 ∃ (s : Finset α), G.IsClique s Finset.map f s = t
      theorem SimpleGraph.IsClique.finsetMap {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {f : α β} {s : Finset α} (h : G.IsClique s) :
      theorem SimpleGraph.IsClique.of_induce {α : Type u_1} {G : SimpleGraph α} {S : G.Subgraph} {F : Set α} {A : Set F} (c : (S.induce F).coe.IsClique A) :

      If a set of vertices A is a clique in subgraph of G induced by a superset of A, its embedding is a clique in G.

      theorem SimpleGraph.IsClique.sdiff_of_sup_edge {α : Type u_1} {G : SimpleGraph α} {v w : α} {s : Set α} (hc : (Gedge v w).IsClique s) :
      G.IsClique (s \ {v})
      theorem SimpleGraph.isClique_sup_edge_of_ne_sdiff {α : Type u_1} {G : SimpleGraph α} {v w : α} {s : Set α} (h : v w) (hv : G.IsClique (s \ {v})) (hw : G.IsClique (s \ {w})) :
      (Gedge v w).IsClique s
      theorem SimpleGraph.isClique_sup_edge_of_ne_iff {α : Type u_1} {G : SimpleGraph α} {v w : α} {s : Set α} (h : v w) :
      (Gedge v w).IsClique s G.IsClique (s \ {v}) G.IsClique (s \ {w})

      n-cliques #

      structure SimpleGraph.IsNClique {α : Type u_1} (G : SimpleGraph α) (n : ) (s : Finset α) :

      An n-clique in a graph is a set of n vertices which are pairwise connected.

      Instances For
        theorem SimpleGraph.isNClique_iff {α : Type u_1} (G : SimpleGraph α) {n : } {s : Finset α} :
        G.IsNClique n s G.IsClique s s.card = n
        @[simp]
        theorem SimpleGraph.isNClique_empty {α : Type u_1} {G : SimpleGraph α} {n : } :
        G.IsNClique n n = 0
        @[simp]
        theorem SimpleGraph.isNClique_singleton {α : Type u_1} {G : SimpleGraph α} {n : } {a : α} :
        G.IsNClique n {a} n = 1
        theorem SimpleGraph.IsNClique.mono {α : Type u_1} {G H : SimpleGraph α} {n : } {s : Finset α} (h : G H) :
        G.IsNClique n sH.IsNClique n s
        theorem SimpleGraph.IsNClique.map {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {n : } {s : Finset α} (h : G.IsNClique n s) {f : α β} :
        theorem SimpleGraph.isNClique_map_iff {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {n : } (hn : 1 < n) {t : Finset β} {f : α β} :
        (SimpleGraph.map f G).IsNClique n t ∃ (s : Finset α), G.IsNClique n s Finset.map f s = t
        @[simp]
        theorem SimpleGraph.isNClique_bot_iff {α : Type u_1} {n : } {s : Finset α} :
        .IsNClique n s n 1 s.card = n
        @[simp]
        theorem SimpleGraph.isNClique_zero {α : Type u_1} {G : SimpleGraph α} {s : Finset α} :
        G.IsNClique 0 s s =
        @[simp]
        theorem SimpleGraph.isNClique_one {α : Type u_1} {G : SimpleGraph α} {s : Finset α} :
        G.IsNClique 1 s ∃ (a : α), s = {a}
        theorem SimpleGraph.IsNClique.insert {α : Type u_1} {G : SimpleGraph α} {n : } {s : Finset α} {a : α} [DecidableEq α] (hs : G.IsNClique n s) (h : bs, G.Adj a b) :
        G.IsNClique (n + 1) (insert a s)
        theorem SimpleGraph.IsNClique.erase_of_mem {α : Type u_1} {G : SimpleGraph α} {n : } {s : Finset α} {a : α} [DecidableEq α] (hs : G.IsNClique n s) (ha : a s) :
        G.IsNClique (n - 1) (s.erase a)
        theorem SimpleGraph.IsNClique.insert_erase {α : Type u_1} {G : SimpleGraph α} {n : } {s : Finset α} {a b : α} [DecidableEq α] (hs : G.IsNClique n s) (ha : ws \ {b}, G.Adj a w) (hb : b s) :
        G.IsNClique n (insert a (s.erase b))
        theorem SimpleGraph.is3Clique_triple_iff {α : Type u_1} {G : SimpleGraph α} {a b c : α} [DecidableEq α] :
        G.IsNClique 3 {a, b, c} G.Adj a b G.Adj a c G.Adj b c
        theorem SimpleGraph.is3Clique_iff {α : Type u_1} {G : SimpleGraph α} {s : Finset α} [DecidableEq α] :
        G.IsNClique 3 s ∃ (a : α) (b : α) (c : α), G.Adj a b G.Adj a c G.Adj b c s = {a, b, c}
        theorem SimpleGraph.is3Clique_iff_exists_cycle_length_three {α : Type u_1} {G : SimpleGraph α} :
        (∃ (s : Finset α), G.IsNClique 3 s) ∃ (u : α) (w : G.Walk u u), w.IsCycle w.length = 3
        theorem SimpleGraph.IsNClique.of_induce {α : Type u_1} {G : SimpleGraph α} {S : G.Subgraph} {F : Set α} {s : Finset { x : α // x F }} {n : } (cc : (S.induce F).coe.IsNClique n s) :
        G.IsNClique n (Finset.map { toFun := Subtype.val, inj' := } s)

        If a set of vertices A is an n-clique in subgraph of G induced by a superset of A, its embedding is an n-clique in G.

        theorem SimpleGraph.IsNClique.erase_of_sup_edge_of_mem {α : Type u_1} {G : SimpleGraph α} [DecidableEq α] {v w : α} {s : Finset α} {n : } (hc : (Gedge v w).IsNClique n s) (hx : v s) :
        G.IsNClique (n - 1) (s.erase v)

        Graphs without cliques #

        def SimpleGraph.CliqueFree {α : Type u_1} (G : SimpleGraph α) (n : ) :

        G.CliqueFree n means that G has no n-cliques.

        Equations
          Instances For
            theorem SimpleGraph.IsNClique.not_cliqueFree {α : Type u_1} {G : SimpleGraph α} {n : } {s : Finset α} (hG : G.IsNClique n s) :
            noncomputable def SimpleGraph.topEmbeddingOfNotCliqueFree {α : Type u_1} {G : SimpleGraph α} {n : } (h : ¬G.CliqueFree n) :

            An embedding of a complete graph that witnesses the fact that the graph is not clique-free.

            Equations
              Instances For
                @[simp]
                @[simp]
                theorem SimpleGraph.cliqueFree_bot {α : Type u_1} {n : } (h : 2 n) :
                theorem SimpleGraph.CliqueFree.mono {α : Type u_1} {G : SimpleGraph α} {m n : } (h : m n) :
                theorem SimpleGraph.CliqueFree.anti {α : Type u_1} {G H : SimpleGraph α} {n : } (h : G H) :
                theorem SimpleGraph.CliqueFree.comap {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {n : } {H : SimpleGraph β} (f : H ↪g G) :

                If a graph is cliquefree, any graph that embeds into it is also cliquefree.

                @[simp]
                theorem SimpleGraph.cliqueFree_map_iff {α : Type u_1} {β : Type u_2} {G : SimpleGraph α} {n : } {f : α β} [Nonempty α] :
                theorem SimpleGraph.cliqueFree_of_card_lt {α : Type u_1} {G : SimpleGraph α} {n : } [Fintype α] (hc : Fintype.card α < n) :

                See SimpleGraph.cliqueFree_of_chromaticNumber_lt for a tighter bound.

                A complete r-partite graph has no n-cliques for r < n.

                def SimpleGraph.completeMultipartiteGraph.topEmbedding {ι : Type u_3} (V : ιType u_4) (f : (i : ι) → V i) :

                Embedding of the complete graph on ι into completeMultipartiteGraph on ι nonempty parts

                Equations
                  Instances For
                    @[simp]
                    theorem SimpleGraph.completeMultipartiteGraph.topEmbedding_apply_snd {ι : Type u_3} (V : ιType u_4) (f : (i : ι) → V i) (i : ι) :
                    ((topEmbedding V f) i).snd = f i
                    @[simp]
                    theorem SimpleGraph.completeMultipartiteGraph.topEmbedding_apply_fst {ι : Type u_3} (V : ιType u_4) (f : (i : ι) → V i) (i : ι) :
                    ((topEmbedding V f) i).fst = i
                    theorem SimpleGraph.completeMultipartiteGraph.not_cliqueFree_of_le_card {n : } {ι : Type u_3} (V : ιType u_4) [Fintype ι] (f : (i : ι) → V i) (hc : n Fintype.card ι) :
                    theorem SimpleGraph.completeMultipartiteGraph.not_cliqueFree_of_le_enatCard {n : } {ι : Type u_3} (V : ιType u_4) (f : (i : ι) → V i) (hc : n ENat.card ι) :
                    theorem SimpleGraph.CliqueFree.replaceVertex {α : Type u_1} {G : SimpleGraph α} {n : } [DecidableEq α] (h : G.CliqueFree n) (s t : α) :

                    Clique-freeness is preserved by replaceVertex.

                    @[simp]
                    theorem SimpleGraph.cliqueFree_one {α : Type u_1} {G : SimpleGraph α} :
                    @[simp]
                    theorem SimpleGraph.cliqueFree_two {α : Type u_1} {G : SimpleGraph α} :
                    theorem SimpleGraph.CliqueFree.mem_of_sup_edge_isNClique {α : Type u_1} {G : SimpleGraph α} {x y : α} {t : Finset α} {n : } (h : G.CliqueFree n) (hc : (Gedge x y).IsNClique n t) :
                    x t
                    theorem SimpleGraph.CliqueFree.sup_edge {α : Type u_1} {G : SimpleGraph α} {n : } (h : G.CliqueFree n) (v w : α) :
                    (Gedge v w).CliqueFree (n + 1)

                    Adding an edge increases the clique number by at most one.

                    theorem SimpleGraph.IsNClique.exists_not_adj_of_cliqueFree_succ {α : Type u_1} {G : SimpleGraph α} {n : } {s : Finset α} (hc : G.IsNClique n s) (h : G.CliqueFree (n + 1)) (x : α) :
                    ys, ¬G.Adj x y
                    theorem SimpleGraph.exists_of_maximal_cliqueFree_not_adj {α : Type u_1} {G : SimpleGraph α} {n : } [DecidableEq α] (h : Maximal (fun (H : SimpleGraph α) => H.CliqueFree (n + 1)) G) {x y : α} (hne : x y) (hn : ¬G.Adj x y) :
                    ∃ (s : Finset α), xs ys G.IsNClique n (insert x s) G.IsNClique n (insert y s)
                    def SimpleGraph.CliqueFreeOn {α : Type u_1} (G : SimpleGraph α) (s : Set α) (n : ) :

                    G.CliqueFreeOn s n means that G has no n-cliques contained in s.

                    Equations
                      Instances For
                        theorem SimpleGraph.CliqueFreeOn.subset {α : Type u_1} (G : SimpleGraph α) {s₁ s₂ : Set α} {n : } (hs : s₁ s₂) (h₂ : G.CliqueFreeOn s₂ n) :
                        G.CliqueFreeOn s₁ n
                        theorem SimpleGraph.CliqueFreeOn.mono {α : Type u_1} (G : SimpleGraph α) {s : Set α} {m n : } (hmn : m n) (hG : G.CliqueFreeOn s m) :
                        theorem SimpleGraph.CliqueFreeOn.anti {α : Type u_1} (G H : SimpleGraph α) {s : Set α} {n : } (hGH : G H) (hH : H.CliqueFreeOn s n) :
                        @[simp]
                        theorem SimpleGraph.cliqueFreeOn_empty {α : Type u_1} (G : SimpleGraph α) {n : } :
                        @[simp]
                        theorem SimpleGraph.cliqueFreeOn_singleton {α : Type u_1} (G : SimpleGraph α) {a : α} {n : } :
                        G.CliqueFreeOn {a} n 1 < n
                        @[simp]
                        theorem SimpleGraph.CliqueFree.cliqueFreeOn {α : Type u_1} (G : SimpleGraph α) {s : Set α} {n : } (hG : G.CliqueFree n) :
                        theorem SimpleGraph.cliqueFreeOn_of_card_lt {α : Type u_1} (G : SimpleGraph α) {n : } {s : Finset α} (h : s.card < n) :
                        G.CliqueFreeOn (↑s) n
                        @[simp]
                        theorem SimpleGraph.cliqueFreeOn_two {α : Type u_1} (G : SimpleGraph α) {s : Set α} :
                        theorem SimpleGraph.CliqueFreeOn.of_succ {α : Type u_1} (G : SimpleGraph α) {s : Set α} {a : α} {n : } (hs : G.CliqueFreeOn s (n + 1)) (ha : a s) :

                        Set of cliques #

                        def SimpleGraph.cliqueSet {α : Type u_1} (G : SimpleGraph α) (n : ) :
                        Set (Finset α)

                        The n-cliques in a graph as a set.

                        Equations
                          Instances For
                            @[simp]
                            theorem SimpleGraph.mem_cliqueSet_iff {α : Type u_1} {G : SimpleGraph α} {n : } {s : Finset α} :
                            @[simp]
                            theorem SimpleGraph.CliqueFree.cliqueSet {α : Type u_1} {G : SimpleGraph α} {n : } :

                            Alias of the reverse direction of SimpleGraph.cliqueSet_eq_empty_iff.

                            theorem SimpleGraph.cliqueSet_mono {α : Type u_1} {G H : SimpleGraph α} {n : } (h : G H) :
                            theorem SimpleGraph.cliqueSet_mono' {α : Type u_1} {G H : SimpleGraph α} (h : G H) :
                            @[simp]
                            theorem SimpleGraph.cliqueSet_zero {α : Type u_1} (G : SimpleGraph α) :
                            @[simp]
                            theorem SimpleGraph.cliqueSet_bot {α : Type u_1} {n : } (hn : 1 < n) :
                            @[simp]
                            theorem SimpleGraph.cliqueSet_map {α : Type u_1} {β : Type u_2} {n : } (hn : n 1) (G : SimpleGraph α) (f : α β) :
                            @[simp]
                            theorem SimpleGraph.cliqueSet_map_of_equiv {α : Type u_1} {β : Type u_2} (G : SimpleGraph α) (e : α β) (n : ) :

                            Clique number #

                            noncomputable def SimpleGraph.cliqueNum {α : Type u_3} (G : SimpleGraph α) :

                            The maximum number of vertices in a clique of a graph G.

                            Equations
                              Instances For
                                theorem SimpleGraph.IsClique.card_le_cliqueNum {α : Type u_3} {G : SimpleGraph α} [Finite α] {t : Finset α} {tc : G.IsClique t} :
                                theorem SimpleGraph.exists_isNClique_cliqueNum {α : Type u_3} {G : SimpleGraph α} [Finite α] :
                                ∃ (s : Finset α), G.IsNClique G.cliqueNum s
                                structure SimpleGraph.IsMaximumClique {α : Type u_3} [Fintype α] (G : SimpleGraph α) (s : Finset α) :

                                A maximum clique in a graph G is a clique with the largest possible size.

                                Instances For
                                  theorem SimpleGraph.isMaximumClique_iff {α : Type u_3} {G : SimpleGraph α} [Fintype α] {s : Finset α} :
                                  G.IsMaximumClique s G.IsClique s ∀ (t : Finset α), G.IsClique tt.card s.card
                                  theorem SimpleGraph.isMaximalClique_iff {α : Type u_3} {G : SimpleGraph α} {s : Set α} :
                                  Maximal G.IsClique s G.IsClique s ∀ (t : Set α), G.IsClique ts tt s

                                  A maximal clique in a graph G is a clique that cannot be extended by adding more vertices.

                                  theorem SimpleGraph.maximumClique_exists {α : Type u_3} {G : SimpleGraph α} [Fintype α] :
                                  ∃ (s : Finset α), G.IsMaximumClique s

                                  Finset of cliques #

                                  def SimpleGraph.cliqueFinset {α : Type u_1} (G : SimpleGraph α) [Fintype α] [DecidableEq α] [DecidableRel G.Adj] (n : ) :

                                  The n-cliques in a graph as a finset.

                                  Equations
                                    Instances For
                                      @[simp]
                                      theorem SimpleGraph.mem_cliqueFinset_iff {α : Type u_1} {G : SimpleGraph α} [Fintype α] [DecidableEq α] [DecidableRel G.Adj] {n : } {s : Finset α} :
                                      @[simp]
                                      theorem SimpleGraph.coe_cliqueFinset {α : Type u_1} (G : SimpleGraph α) [Fintype α] [DecidableEq α] [DecidableRel G.Adj] (n : ) :
                                      (G.cliqueFinset n) = G.cliqueSet n

                                      Alias of the reverse direction of SimpleGraph.cliqueFinset_eq_empty_iff.

                                      @[simp]
                                      theorem SimpleGraph.cliqueFinset_map {α : Type u_1} {β : Type u_2} (G : SimpleGraph α) [Fintype α] [DecidableEq α] [DecidableRel G.Adj] {n : } [Fintype β] [DecidableEq β] (f : α β) (hn : n 1) :
                                      (SimpleGraph.map f G).cliqueFinset n = Finset.map { toFun := Finset.map f, inj' := } (G.cliqueFinset n)
                                      @[simp]
                                      theorem SimpleGraph.cliqueFinset_map_of_equiv {α : Type u_1} {β : Type u_2} (G : SimpleGraph α) [Fintype α] [DecidableEq α] [DecidableRel G.Adj] [Fintype β] [DecidableEq β] (e : α β) (n : ) :

                                      Independent Sets #

                                      @[reducible, inline]
                                      abbrev SimpleGraph.IsIndepSet {α : Type u_1} (G : SimpleGraph α) (s : Set α) :

                                      An independent set in a graph is a set of vertices that are pairwise not adjacent.

                                      Equations
                                        Instances For
                                          theorem SimpleGraph.isIndepSet_iff {α : Type u_1} (G : SimpleGraph α) {s : Set α} :
                                          G.IsIndepSet s s.Pairwise fun (v w : α) => ¬G.Adj v w
                                          @[simp]
                                          theorem SimpleGraph.isClique_compl {α : Type u_1} (G : SimpleGraph α) {s : Set α} :

                                          An independent set is a clique in the complement graph and vice versa.

                                          @[simp]
                                          theorem SimpleGraph.isIndepSet_compl {α : Type u_1} (G : SimpleGraph α) {s : Set α} :

                                          An independent set in the complement graph is a clique and vice versa.

                                          theorem SimpleGraph.IsIndepSet.nonempty_mem_compl_mem_edge {α : Type u_1} (G : SimpleGraph α) [Fintype α] [DecidableEq α] {s : Finset α} (indA : G.IsIndepSet s) {e : Sym2 α} (he : e G.edgeSet) :
                                          {bs | b e}.Nonempty

                                          If s is an independent set, its complement meets every edge of G.

                                          The neighbors of a vertex v form an independent set in a triangle free graph G.

                                          theorem SimpleGraph.isIndepSet_induce {α : Type u_1} (G : SimpleGraph α) {F : Set α} {s : Set F} :

                                          The embedding of an independent set of an induced subgraph of the subgraph G is an independent set in G and vice versa.

                                          N-Independent sets #

                                          structure SimpleGraph.IsNIndepSet {α : Type u_1} (G : SimpleGraph α) (n : ) (s : Finset α) :

                                          An n-independent set in a graph is a set of n vertices which are pairwise nonadjacent.

                                          Instances For
                                            theorem SimpleGraph.isNIndepSet_iff {α : Type u_1} (G : SimpleGraph α) (n : ) (s : Finset α) :
                                            G.IsNIndepSet n s G.IsIndepSet s s.card = n
                                            @[simp]
                                            theorem SimpleGraph.isNClique_compl {α : Type u_1} (G : SimpleGraph α) {n : } {s : Finset α} :

                                            An n-independent set is an n-clique in the complement graph and vice versa.

                                            @[simp]
                                            theorem SimpleGraph.isNIndepSet_compl {α : Type u_1} (G : SimpleGraph α) {n : } {s : Finset α} :

                                            An n-independent set in the complement graph is an n-clique and vice versa.

                                            theorem SimpleGraph.isNIndepSet_induce {α : Type u_1} (G : SimpleGraph α) {F : Set α} {s : Finset { x : α // x F }} {n : } :
                                            (.induce F).coe.IsNIndepSet n s G.IsNIndepSet n (Finset.map { toFun := Subtype.val, inj' := } s)

                                            The embedding of an n-independent set of an induced subgraph of the subgraph G is an n-independent set in G and vice versa.

                                            Graphs without independent sets #

                                            def SimpleGraph.IndepSetFree {α : Type u_1} (G : SimpleGraph α) (n : ) :

                                            G.IndepSetFree n means that G has no n-independent sets.

                                            Equations
                                              Instances For
                                                @[simp]

                                                An graph is n-independent set free iff its complement is n-clique free.

                                                @[simp]

                                                An graph's complement is n-independent set free iff it is n-clique free.

                                                def SimpleGraph.IndepSetFreeOn {α : Type u_1} (G : SimpleGraph α) (s : Set α) (n : ) :

                                                G.IndepSetFreeOn s n means that G has no n-independent sets contained in s.

                                                Equations
                                                  Instances For

                                                    Set of independent sets #

                                                    def SimpleGraph.indepSetSet {α : Type u_1} (G : SimpleGraph α) (n : ) :
                                                    Set (Finset α)

                                                    The n-independent sets in a graph as a set.

                                                    Equations
                                                      Instances For
                                                        @[simp]
                                                        theorem SimpleGraph.mem_indepSetSet_iff {α : Type u_1} {G : SimpleGraph α} {n : } {s : Finset α} :

                                                        Independence Number #

                                                        noncomputable def SimpleGraph.indepNum {α : Type u_3} (G : SimpleGraph α) :

                                                        The maximal number of vertices of an independent set in a graph G.

                                                        Equations
                                                          Instances For
                                                            @[simp]
                                                            @[simp]
                                                            theorem SimpleGraph.IsIndepSet.card_le_indepNum {α : Type u_3} {G : SimpleGraph α} [Finite α] {t : Finset α} (tc : G.IsIndepSet t) :
                                                            structure SimpleGraph.IsMaximumIndepSet {α : Type u_3} [Fintype α] (G : SimpleGraph α) (s : Finset α) :

                                                            An independent set in a graph G such that there is no independent set with more vertices.

                                                            Instances For
                                                              theorem SimpleGraph.isMaximumIndepSet_iff {α : Type u_3} [Fintype α] (G : SimpleGraph α) (s : Finset α) :
                                                              G.IsMaximumIndepSet s G.IsIndepSet s ∀ (t : Finset α), G.IsIndepSet tt.card s.card
                                                              theorem SimpleGraph.isMaximalIndepSet_iff {α : Type u_3} {G : SimpleGraph α} {s : Set α} :
                                                              Maximal G.IsIndepSet s G.IsIndepSet s ∀ (t : Set α), G.IsIndepSet ts tt s

                                                              An independent set in a graph G that cannot be extended by adding more vertices.

                                                              @[simp]
                                                              @[simp]
                                                              theorem SimpleGraph.maximumIndepSet_exists {α : Type u_3} {G : SimpleGraph α} [Fintype α] :
                                                              ∃ (s : Finset α), G.IsMaximumIndepSet s

                                                              Finset of independent sets #

                                                              The n-independent sets in a graph as a finset.

                                                              Equations
                                                                Instances For
                                                                  @[simp]
                                                                  theorem SimpleGraph.mem_indepSetFinset_iff {α : Type u_1} {G : SimpleGraph α} [Fintype α] [DecidableEq α] [DecidableRel G.Adj] {n : } {s : Finset α} :