Basic properties of Azumaya algebras #
In this file we prove basic facts about Azumaya algebras such as R
is an Azumaya algebra
over itself where R
is a commutative ring.
Main Results #
IsAzumaya.id
:R
is an Azumaya algebra over itself.IsAzumaya.ofAlgEquiv
: IfA
is an Azumaya algebra overR
andA
is isomorphic toB
as anR
-algebra, thenB
is an Azumaya algebra overR
.
Tags #
Noncommutative algebra, Azumaya algebra, Brauer Group
theorem
IsAzumaya.AlgHom.mulLeftRight_bij
(R : Type u_1)
(A : Type u_2)
[CommSemiring R]
[Ring A]
[Algebra R A]
[h : IsAzumaya R A]
:
@[reducible, inline]
The "canonical" isomorphism between R ⊗ Rᵒᵖ
and End R R
which is equal
to AlgHom.mulLeftRight R R
.
Equations
Instances For
theorem
IsAzumaya.mulLeftRight_comp_congr
(R : Type u_1)
(A : Type u_2)
(B : Type u_3)
[CommSemiring R]
[Ring A]
[Ring B]
[Algebra R A]
[Algebra R B]
(e : A ≃ₐ[R] B)
:
(AlgHom.mulLeftRight R B).comp ↑(Algebra.TensorProduct.congr e (AlgEquiv.op e)) = (↑e.toLinearEquiv.algConj).comp (AlgHom.mulLeftRight R A)
The following diagram commutes:
e ⊗ eᵒᵖ
A ⊗ Aᵐᵒᵖ ------------> B ⊗ Bᵐᵒᵖ
| |
| |
| mulLeftRight R A | mulLeftRight R B
| |
V V
End R A ------------> End R B
e.conj