Skip to content

Commit

Permalink
Beef up ring homomorphism to an algebra homomorphism task (#218)
Browse files Browse the repository at this point in the history
* First attempt

* Finished attempt

* Formatting

* Finalised for PR

* Update FLT/DedekindDomain/FiniteAdeleRing/BaseChange.lean

Co-authored-by: Kevin Buzzard <[email protected]>

* Update FLT/DedekindDomain/FiniteAdeleRing/BaseChange.lean

Co-authored-by: Kevin Buzzard <[email protected]>

* Update FLT/DedekindDomain/FiniteAdeleRing/BaseChange.lean

Co-authored-by: Kevin Buzzard <[email protected]>

* Update FLT/DedekindDomain/FiniteAdeleRing/BaseChange.lean

Co-authored-by: Kevin Buzzard <[email protected]>

---------

Co-authored-by: Kevin Buzzard <[email protected]>
  • Loading branch information
WilliamCoram and kbuzzard authored Nov 17, 2024
1 parent bbee48f commit afa70b6
Showing 1 changed file with 16 additions and 5 deletions.
21 changes: 16 additions & 5 deletions FLT/DedekindDomain/FiniteAdeleRing/BaseChange.lean
Original file line number Diff line number Diff line change
Expand Up @@ -169,13 +169,24 @@ noncomputable local instance (w : HeightOneSpectrum B) :
Algebra K (adicCompletion L w) := RingHom.toAlgebra <|
(algebraMap L (adicCompletion L w)).comp (algebraMap K L)


variable {B L} in
noncomputable def adicCompletion_comap_algHom (w : HeightOneSpectrum B) :
letI : Algebra K (adicCompletion L w) := RingHom.toAlgebra <|
(algebraMap L (adicCompletion L w)).comp (algebraMap K L);
letI : Module K (adicCompletion L w) := Algebra.toModule
(HeightOneSpectrum.adicCompletion K (comap A w)) →ₐ[K] (HeightOneSpectrum.adicCompletion L w) :=
sorry -- use `adicCompletion_comap_ringHom` and prove it's a K-algebra homomorphism
(HeightOneSpectrum.adicCompletion K (comap A w)) →ₐ[K]
(HeightOneSpectrum.adicCompletion L w) where
__ := adicCompletion_comap_ringHom A K w
commutes' r := by
simp only [RingHom.toMonoidHom_eq_coe, OneHom.toFun_eq_coe, MonoidHom.toOneHom_coe,
MonoidHom.coe_coe]
have : (adicCompletion_comap_ringHom A K w) (r : adicCompletion K (comap A w)) =
(algebraMap L (adicCompletion L w)) (algebraMap K L r) := by
letI : UniformSpace L := w.adicValued.toUniformSpace
letI : UniformSpace K := (comap A w).adicValued.toUniformSpace
rw [adicCompletion_comap_ringHom, UniformSpace.Completion.mapRingHom]
rw [show (r : adicCompletion K (comap A w)) = @UniformSpace.Completion.coe' K this r from rfl]
apply UniformSpace.Completion.extensionHom_coe
rw [this]
rfl

end IsDedekindDomain.HeightOneSpectrum

Expand Down

0 comments on commit afa70b6

Please sign in to comment.