Skip to content

Commit

Permalink
move MvPolynomial.transcendental_X to earlier file
Browse files Browse the repository at this point in the history
  • Loading branch information
acmepjz committed Oct 31, 2024
1 parent c3166bf commit 61e165b
Show file tree
Hide file tree
Showing 2 changed files with 40 additions and 4 deletions.
40 changes: 40 additions & 0 deletions Mathlib/RingTheory/Algebraic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,7 @@ import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic
import Mathlib.RingTheory.Polynomial.IntegralNormalization
import Mathlib.RingTheory.LocalRing.Basic
import Mathlib.Algebra.Polynomial.Lifts
import Mathlib.Algebra.MvPolynomial.Supported

/-!
# Algebraic elements and algebraic extensions
Expand Down Expand Up @@ -697,3 +698,42 @@ theorem Polynomial.algebraMap_pi_self_eq_eval :
rfl

end Pi

namespace MvPolynomial

variable {σ : Type*} (R : Type*) [CommRing R]

theorem transcendental_supported_X {i : σ} {s : Set σ} (h : i ∉ s) :
Transcendental (supported R s) (X i : MvPolynomial σ R) := by
rw [transcendental_iff_injective]
let f := (Subalgebra.val _).comp ((optionEquivLeft R s).symm |>.trans
(renameEquiv R (Set.subtypeInsertEquivOption h)).symm |>.trans
(supportedEquivMvPolynomial _).symm).toAlgHom
let g := ((Polynomial.aeval (R := supported R s) (X i : MvPolynomial σ R)).restrictScalars R).comp
(Polynomial.mapAlgEquiv (supportedEquivMvPolynomial s).symm).toAlgHom
have hinj : Function.Injective f := by
simp only [f, AlgEquiv.toAlgHom_eq_coe, AlgHom.coe_comp, AlgHom.coe_coe, Subalgebra.coe_val,
EquivLike.injective_comp, Subtype.val_injective]
have hfg : f = g := by
ext1
· ext1
simp [f, g, Set.subtypeInsertEquivOption, Subalgebra.algebraMap_eq]
· simp [f, g, Set.subtypeInsertEquivOption]
simpa only [hfg, g, AlgEquiv.toAlgHom_eq_coe, AlgHom.coe_comp, AlgHom.coe_coe,
EquivLike.injective_comp, AlgHom.coe_restrictScalars'] using hinj

theorem transcendental_X (i : σ) : Transcendental R (X i : MvPolynomial σ R) := by
have := transcendental_supported_X R (Set.not_mem_empty i)
let f := (Algebra.botEquivOfInjective (MvPolynomial.C_injective σ R)).symm.trans
(Subalgebra.equivOfEq _ _ supported_empty).symm |>.toRingEquiv
rwa [Transcendental, ← isAlgebraic_ringHom_iff f (RingHom.id (MvPolynomial σ R))
Function.injective_id (by ext1; rfl), RingHom.id_apply, ← Transcendental]

theorem transcendental_supported_X_iff [Nontrivial R] {i : σ} {s : Set σ} :
Transcendental (supported R s) (X i : MvPolynomial σ R) ↔ i ∉ s := by
refine ⟨?_, transcendental_supported_X R⟩
rw [Transcendental, not_imp_not]
exact fun h ↦ isAlgebraic_algebraMap
(⟨X i, Algebra.subset_adjoin (Set.mem_image_of_mem _ h)⟩ : supported R s)

end MvPolynomial
4 changes: 0 additions & 4 deletions Mathlib/RingTheory/AlgebraicIndependent.lean
Original file line number Diff line number Diff line change
Expand Up @@ -173,10 +173,6 @@ theorem algebraicIndependent_X : AlgebraicIndependent R (X (R := R) (σ := σ))
rw [AlgebraicIndependent, aeval_X_left]
exact injective_id

variable {σ} in
theorem transcendental_X (i : σ) : Transcendental R (X (R := R) i) :=
(algebraicIndependent_X σ R).transcendental i

end MvPolynomial

open AlgebraicIndependent
Expand Down

0 comments on commit 61e165b

Please sign in to comment.