Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - perf: speed up galActionHom_bijective_of_prime_degree #10608

Closed
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
7 changes: 3 additions & 4 deletions Mathlib/Analysis/Complex/Polynomial.lean
Original file line number Diff line number Diff line change
Expand Up @@ -128,8 +128,6 @@ theorem galActionHom_bijective_of_prime_degree {p : ℚ[X]} (p_irr : Irreducible
rw [Multiset.toFinset_card_of_nodup, ← natDegree_eq_card_roots]
· exact IsAlgClosed.splits_codomain p
· exact nodup_roots ((separable_map (algebraMap ℚ ℂ)).mpr p_irr.separable)
have h2 : Fintype.card p.Gal = Fintype.card (galActionHom p ℂ).range :=
Fintype.card_congr (MonoidHom.ofInjective (galActionHom_injective p ℂ)).toEquiv
let conj' := restrict p ℂ (Complex.conjAe.restrictScalars ℚ)
refine'
⟨galActionHom_injective p ℂ, fun x =>
Expand All @@ -138,8 +136,9 @@ theorem galActionHom_bijective_of_prime_degree {p : ℚ[X]} (p_irr : Irreducible
apply Equiv.Perm.subgroup_eq_top_of_swap_mem
· rwa [h1]
· rw [h1]
convert prime_degree_dvd_card p_irr p_deg using 1
convert h2.symm
simpa only [Fintype.card_eq_nat_card,
Nat.card_congr (MonoidHom.ofInjective (galActionHom_injective p ℂ)).toEquiv.symm]
using prime_degree_dvd_card p_irr p_deg
· exact ⟨conj', rfl⟩
· rw [← Equiv.Perm.card_support_eq_two]
apply Nat.add_left_cancel
Expand Down
5 changes: 5 additions & 0 deletions Mathlib/SetTheory/Cardinal/Finite.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,6 +42,11 @@ theorem card_eq_fintype_card [Fintype α] : Nat.card α = Fintype.card α :=
mk_toNat_eq_card
#align nat.card_eq_fintype_card Nat.card_eq_fintype_card

/-- Because this theorem takes `Fintype α` as a non-instance argument, it can be used in particular
when `Fintype.card` ends up with different instance than the one found by inference -/
theorem _root_.Fintype.card_eq_nat_card {_ : Fintype α} : Fintype.card α = Nat.card α :=
mk_toNat_eq_card.symm

lemma card_eq_finsetCard (s : Finset α) : Nat.card s = s.card := by
simp only [Nat.card_eq_fintype_card, Fintype.card_coe]

Expand Down
Loading