Skip to content

Commit

Permalink
added a theorem on existence of a maximum clique
Browse files Browse the repository at this point in the history
  • Loading branch information
ooovi committed Oct 31, 2024
1 parent a2cfddb commit a3e14e0
Showing 1 changed file with 12 additions and 3 deletions.
15 changes: 12 additions & 3 deletions Mathlib/Combinatorics/SimpleGraph/Clique.lean
Original file line number Diff line number Diff line change
Expand Up @@ -635,15 +635,24 @@ lemma fintype_cliqueNum_bddAbove : BddAbove {n | ∃ s, G.IsNClique n s} := by
theorem clique_card_le_cliqueNum (t : Finset α) (tc : G.IsClique t) : #t ≤ G.cliqueNum :=
le_csSup fintype_cliqueNum_bddAbove (Exists.intro t ⟨tc, rfl⟩)

lemma cliqueNum_attained : G.cliqueNum ∈ {n | ∃ s, G.IsNClique n s} :=
Nat.sSup_mem ⟨0, by simp[isNClique_empty.mpr rfl]⟩ fintype_cliqueNum_bddAbove

theorem maximumClique_card_eq_cliqueNum (t : Finset α) (tmc : G.isMaximumClique t) :
#t = G.cliqueNum := by
let ⟨tclique, tmax⟩ := tmc
refine eq_of_le_of_not_lt (clique_card_le_cliqueNum _ tclique) ?_
have ⟨s, sclique, scn⟩ : G.cliqueNum ∈ {n | ∃ s, G.IsNClique n s} :=
Nat.sSup_mem ⟨ 0, by simp[isNClique_empty.mpr rfl] ⟩ fintype_cliqueNum_bddAbove
rw [cliqueNum, ← scn]
have ⟨s, sclique, scn⟩ := G.cliqueNum_attained
rw [← scn]
exact LE.le.not_lt (tmax s sclique)

theorem maximumClique_exists : ∃ s , G.isMaximumClique s := by
have ⟨s, smax⟩ := G.cliqueNum_attained
simp_all [isMaximumClique]
refine Exists.intro s ⟨smax.clique , ?_⟩
rw [smax.card_eq]
exact clique_card_le_cliqueNum

end CliqueNumber

end SimpleGraph

0 comments on commit a3e14e0

Please sign in to comment.