Skip to content

Commit 44b647d

Browse files
committed
feat(FieldTheory/PolynomialGaloisGroup): The Galois group of an irreducible polynomial acts transitively on the roots (#9121)
This PR adds a lemma stating that the Galois group of an irreducible polynomial acts transitively on the roots.
1 parent a8d482d commit 44b647d

File tree

1 file changed

+8
-0
lines changed

1 file changed

+8
-0
lines changed

Mathlib/FieldTheory/PolynomialGaloisGroup.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -198,6 +198,14 @@ instance galAction [Fact (p.Splits (algebraMap F E))] : MulAction p.Gal (rootSet
198198
simp only [smul_def, Equiv.apply_symm_apply, Equiv.symm_apply_apply, mul_smul]
199199
#align polynomial.gal.gal_action Polynomial.Gal.galAction
200200

201+
lemma galAction_isPretransitive [Fact (p.Splits (algebraMap F E))] (hp : Irreducible p) :
202+
MulAction.IsPretransitive p.Gal (p.rootSet E) := by
203+
refine' ⟨fun x y ↦ _⟩
204+
have hx := minpoly.eq_of_irreducible hp (mem_rootSet.mp ((rootsEquivRoots p E).symm x).2).2
205+
have hy := minpoly.eq_of_irreducible hp (mem_rootSet.mp ((rootsEquivRoots p E).symm y).2).2
206+
obtain ⟨g, hg⟩ := (Normal.minpoly_eq_iff_mem_orbit p.SplittingField).mp (hy.symm.trans hx)
207+
exact ⟨g, (rootsEquivRoots p E).apply_eq_iff_eq_symm_apply.mpr (Subtype.ext hg)⟩
208+
201209
variable {p E}
202210

203211
/-- `Polynomial.Gal.restrict p E` is compatible with `Polynomial.Gal.galAction p E`. -/

0 commit comments

Comments
 (0)