Skip to content

Commit

Permalink
feat: Compare card of subgroup to card of group (#5347)
Browse files Browse the repository at this point in the history
Includes new lemmas, one that shows that the cardinality of a subgroup is at most the cardinality of the ambient group, and others which shows that the cardinality of the top group is equal to that of the ambient group, and that in fact this is iff.

Co-authored-by: Floris van Doorn
  • Loading branch information
BoltonBailey committed Jul 3, 2023
1 parent 4ebd604 commit 5cd26e5
Showing 1 changed file with 13 additions and 0 deletions.
13 changes: 13 additions & 0 deletions Mathlib/GroupTheory/Subgroup/Finite.lean
Expand Up @@ -128,6 +128,11 @@ theorem card_bot {_ : Fintype (⊥ : Subgroup G)} : Fintype.card (⊥ : Subgroup
#align subgroup.card_bot Subgroup.card_bot
#align add_subgroup.card_bot AddSubgroup.card_bot

@[to_additive]
theorem card_top [Fintype G] : Fintype.card (⊤ : Subgroup G) = Fintype.card G := by
rw [Fintype.card_eq]
exact Nonempty.intro Subgroup.topEquiv.toEquiv

@[to_additive]
theorem eq_top_of_card_eq [Fintype H] [Fintype G] (h : Fintype.card H = Fintype.card G) :
H = ⊤ := by
Expand All @@ -138,6 +143,10 @@ theorem eq_top_of_card_eq [Fintype H] [Fintype G] (h : Fintype.card H = Fintype.
#align subgroup.eq_top_of_card_eq Subgroup.eq_top_of_card_eq
#align add_subgroup.eq_top_of_card_eq AddSubgroup.eq_top_of_card_eq

@[to_additive (attr := simp)]
theorem card_eq_iff_eq_top [Fintype H] [Fintype G] : Fintype.card H = Fintype.card G ↔ H = ⊤ :=
Iff.intro (eq_top_of_card_eq H) (fun h ↦ by simpa only [h] using card_top)

@[to_additive]
theorem eq_top_of_le_card [Fintype H] [Fintype G] (h : Fintype.card G ≤ Fintype.card H) : H = ⊤ :=
eq_top_of_card_eq H
Expand Down Expand Up @@ -173,6 +182,10 @@ theorem one_lt_card_iff_ne_bot [Fintype H] : 1 < Fintype.card H ↔ H ≠ ⊥ :=
#align subgroup.one_lt_card_iff_ne_bot Subgroup.one_lt_card_iff_ne_bot
#align add_subgroup.pos_card_iff_ne_bot AddSubgroup.one_lt_card_iff_ne_bot

@[to_additive]
theorem card_le_card_group [Fintype G] [Fintype H] : Fintype.card H ≤ Fintype.card G :=
Fintype.card_le_of_injective _ Subtype.coe_injective

end Subgroup

namespace Subgroup
Expand Down

0 comments on commit 5cd26e5

Please sign in to comment.