Skip to content

Commit

Permalink
feat: generalize Algebra.TensorProduct.rightAlgebra (#18089)
Browse files Browse the repository at this point in the history
Co-authored-by: Kevin Buzzard <[email protected]>
  • Loading branch information
Ruben-VandeVelde and kbuzzard committed Oct 23, 2024
1 parent 11c09b7 commit 80a505f
Showing 1 changed file with 11 additions and 3 deletions.
14 changes: 11 additions & 3 deletions Mathlib/RingTheory/TensorProduct/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -576,12 +576,22 @@ instance instCommRing : CommRing (A ⊗[R] B) :=
{ toRing := inferInstance
mul_comm := mul_comm }

end CommRing

section RightAlgebra

variable [CommSemiring R]
variable [Semiring A] [Algebra R A]
variable [CommSemiring B] [Algebra R B]

/-- `S ⊗[R] T` has a `T`-algebra structure. This is not a global instance or else the action of
`S` on `S ⊗[R] S` would be ambiguous. -/
abbrev rightAlgebra : Algebra B (A ⊗[R] B) :=
(Algebra.TensorProduct.includeRight.toRingHom : B →+* A ⊗[R] B).toAlgebra
includeRight.toRingHom.toAlgebra' fun b x => by
suffices LinearMap.mulLeft R (includeRight b) = LinearMap.mulRight R (includeRight b) from
congr($this x)
ext xa xb
simp [mul_comm]

attribute [local instance] TensorProduct.rightAlgebra

Expand All @@ -590,8 +600,6 @@ instance right_isScalarTower : IsScalarTower R B (A ⊗[R] B) :=

end RightAlgebra

end CommRing

/-- Verify that typeclass search finds the ring structure on `A ⊗[ℤ] B`
when `A` and `B` are merely rings, by treating both as `ℤ`-algebras.
-/
Expand Down

0 comments on commit 80a505f

Please sign in to comment.