Quadratic Algebra #
In this file we define the quadratic algebra QuadraticAlgebra R a b
over a commutative ring R
,
and define some algebraic structures on it.
Main definitions #
QuadraticAlgebra R a b
: Bourbaki, Algebra I with coefficientsa
,b
inR
.
Tags #
Quadratic algebra, quadratic extension
Quadratic algebra over a type with fixed coefficient where $i^2 = a + bi$, implemented as
a structure with two fields, re
and im
. When R
is a commutative ring, this is isomorphic to
R[X]/(X^2-b*X-a)
.
- re : R
Real part of an element in quadratic algebra
- im : R
Imaginary part of an element in quadratic algebra
Instances For
instance
instDecidableEqQuadraticAlgebra
{R✝ : Type u_1}
{a✝ b✝ : R✝}
[DecidableEq R✝]
:
DecidableEq (QuadraticAlgebra R✝ a✝ b✝)
The equivalence between quadratic algebra over R
and R × R
.
Equations
Instances For
@[simp]
instance
QuadraticAlgebra.instSubsingleton
{R : Type u_1}
{a b : R}
[Subsingleton R]
:
Subsingleton (QuadraticAlgebra R a b)
instance
QuadraticAlgebra.instNontrivial
{R : Type u_1}
{a b : R}
[Nontrivial R]
:
Nontrivial (QuadraticAlgebra R a b)
Coercion R → QuadraticAlgebra R a b
.
Instances For
Equations
instance
QuadraticAlgebra.instInhabited
{R : Type u_1}
{a b : R}
[Zero R]
:
Inhabited (QuadraticAlgebra R a b)
Equations
- QuadraticAlgebra.instInhabited = { default := 0 }
instance
QuadraticAlgebra.instOne
{R : Type u_1}
{a b : R}
[Zero R]
[One R]
:
One (QuadraticAlgebra R a b)
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
instance
QuadraticAlgebra.instSMul
{R : Type u_1}
{S : Type u_2}
{a b : R}
[SMul S R]
:
SMul S (QuadraticAlgebra R a b)
instance
QuadraticAlgebra.instIsScalarTower
{R : Type u_1}
{S : Type u_2}
{T : Type u_3}
{a b : R}
[SMul S R]
[SMul T R]
[SMul S T]
[IsScalarTower S T R]
:
IsScalarTower S T (QuadraticAlgebra R a b)
instance
QuadraticAlgebra.instSMulCommClass
{R : Type u_1}
{S : Type u_2}
{T : Type u_3}
{a b : R}
[SMul S R]
[SMul T R]
[SMulCommClass S T R]
:
SMulCommClass S T (QuadraticAlgebra R a b)
instance
QuadraticAlgebra.instIsCentralScalar
{R : Type u_1}
{S : Type u_2}
{a b : R}
[SMul S R]
[SMul Sᵐᵒᵖ R]
[IsCentralScalar S R]
:
IsCentralScalar S (QuadraticAlgebra R a b)
instance
QuadraticAlgebra.instMulAction
{R : Type u_1}
{S : Type u_2}
{a b : R}
[Monoid S]
[MulAction S R]
:
MulAction S (QuadraticAlgebra R a b)
Equations
- QuadraticAlgebra.instMulAction = { toSMul := QuadraticAlgebra.instSMul, one_smul := ⋯, mul_smul := ⋯ }
@[simp]
theorem
QuadraticAlgebra.coe_smul
{R : Type u_1}
{S : Type u_2}
{a b : R}
[Zero R]
[SMulZeroClass S R]
(s : S)
(r : R)
:
instance
QuadraticAlgebra.instAddMonoid
{R : Type u_1}
{a b : R}
[AddMonoid R]
:
AddMonoid (QuadraticAlgebra R a b)
Equations
- One or more equations did not get rendered due to their size.
instance
QuadraticAlgebra.instDistribMulAction
{R : Type u_1}
{S : Type u_2}
{a b : R}
[Monoid S]
[AddMonoid R]
[DistribMulAction S R]
:
DistribMulAction S (QuadraticAlgebra R a b)
Equations
- QuadraticAlgebra.instDistribMulAction = { toMulAction := QuadraticAlgebra.instMulAction, smul_zero := ⋯, smul_add := ⋯ }
instance
QuadraticAlgebra.instAddCommMonoid
{R : Type u_1}
{a b : R}
[AddCommMonoid R]
:
AddCommMonoid (QuadraticAlgebra R a b)
Equations
- QuadraticAlgebra.instAddCommMonoid = { toAddMonoid := QuadraticAlgebra.instAddMonoid, add_comm := ⋯ }
instance
QuadraticAlgebra.instModule
{R : Type u_1}
{S : Type u_2}
{a b : R}
[Semiring S]
[AddCommMonoid R]
[Module S R]
:
Module S (QuadraticAlgebra R a b)
Equations
- QuadraticAlgebra.instModule = { toDistribMulAction := QuadraticAlgebra.instDistribMulAction, add_smul := ⋯, zero_smul := ⋯ }
instance
QuadraticAlgebra.instAddGroup
{R : Type u_1}
{a b : R}
[AddGroup R]
:
AddGroup (QuadraticAlgebra R a b)
Equations
- One or more equations did not get rendered due to their size.
instance
QuadraticAlgebra.instAddCommGroup
{R : Type u_1}
{a b : R}
[AddCommGroup R]
:
AddCommGroup (QuadraticAlgebra R a b)
Equations
- QuadraticAlgebra.instAddCommGroup = { toAddGroup := QuadraticAlgebra.instAddGroup, add_comm := ⋯ }
instance
QuadraticAlgebra.instAddCommMonoidWithOne
{R : Type u_1}
{a b : R}
[AddCommMonoidWithOne R]
:
AddCommMonoidWithOne (QuadraticAlgebra R a b)
Equations
- One or more equations did not get rendered due to their size.
@[simp]
theorem
QuadraticAlgebra.coe_ofNat
{R : Type u_1}
{a b : R}
[AddCommMonoidWithOne R]
(n : ℕ)
[n.AtLeastTwo]
:
@[simp]
@[simp]
theorem
QuadraticAlgebra.re_ofNat
{R : Type u_1}
{a b : R}
[AddCommMonoidWithOne R]
(n : ℕ)
[n.AtLeastTwo]
:
theorem
QuadraticAlgebra.im_ofNat
{R : Type u_1}
{a b : R}
[AddCommMonoidWithOne R]
(n : ℕ)
[n.AtLeastTwo]
:
instance
QuadraticAlgebra.instAddCommGroupWithOne
{R : Type u_1}
{a b : R}
[AddCommGroupWithOne R]
:
AddCommGroupWithOne (QuadraticAlgebra R a b)
Equations
- One or more equations did not get rendered due to their size.
@[simp]
@[simp]
instance
QuadraticAlgebra.instNonUnitalNonAssocSemiring
{R : Type u_1}
{a b : R}
[NonUnitalNonAssocSemiring R]
:
Equations
- One or more equations did not get rendered due to their size.
theorem
QuadraticAlgebra.coe_mul_eq_smul
{R : Type u_1}
{a b : R}
[NonUnitalNonAssocSemiring R]
(r : R)
(x : QuadraticAlgebra R a b)
:
@[simp]
instance
QuadraticAlgebra.instNonAssocSemiring
{R : Type u_1}
{a b : R}
[NonAssocSemiring R]
:
NonAssocSemiring (QuadraticAlgebra R a b)
Equations
- One or more equations did not get rendered due to their size.
QuadraticAlgebra.re
as a LinearMap
Equations
- QuadraticAlgebra.reₗ a b = { toFun := QuadraticAlgebra.re, map_add' := ⋯, map_smul' := ⋯ }
Instances For
@[simp]
theorem
QuadraticAlgebra.reₗ_apply
{R : Type u_1}
(a b : R)
[Semiring R]
(self : QuadraticAlgebra R a b)
:
QuadraticAlgebra.im
as a LinearMap
Equations
- QuadraticAlgebra.imₗ a b = { toFun := QuadraticAlgebra.im, map_add' := ⋯, map_smul' := ⋯ }
Instances For
@[simp]
theorem
QuadraticAlgebra.imₗ_apply
{R : Type u_1}
(a b : R)
[Semiring R]
(self : QuadraticAlgebra R a b)
:
QuadraticAlgebra.equivTuple
as a LinearEquiv
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
theorem
QuadraticAlgebra.linearEquivTuple_apply
{R : Type u_1}
(a b : R)
[Semiring R]
(z : QuadraticAlgebra R a b)
:
@[simp]
theorem
QuadraticAlgebra.linearEquivTuple_symm_apply
{R : Type u_1}
(a b : R)
[Semiring R]
(x : Fin 2 → R)
:
noncomputable def
QuadraticAlgebra.basis
{R : Type u_1}
(a b : R)
[Semiring R]
:
Module.Basis (Fin 2) R (QuadraticAlgebra R a b)
QuadraticAlgebra R a b
has a basis over R
given by 1
and i
Equations
Instances For
instance
QuadraticAlgebra.instFinite
{R : Type u_1}
(a b : R)
[Semiring R]
:
Module.Finite R (QuadraticAlgebra R a b)
instance
QuadraticAlgebra.instFree
{R : Type u_1}
(a b : R)
[Semiring R]
:
Module.Free R (QuadraticAlgebra R a b)
theorem
QuadraticAlgebra.rank_eq_two
{R : Type u_1}
(a b : R)
[Semiring R]
[StrongRankCondition R]
:
theorem
QuadraticAlgebra.finrank_eq_two
{R : Type u_1}
(a b : R)
[Semiring R]
[StrongRankCondition R]
:
instance
QuadraticAlgebra.instCommSemiring
{R : Type u_1}
{a b : R}
[CommSemiring R]
:
CommSemiring (QuadraticAlgebra R a b)
Equations
- One or more equations did not get rendered due to their size.
instance
QuadraticAlgebra.instAlgebra
{R : Type u_1}
{S : Type u_2}
{a b : R}
[CommSemiring S]
[CommSemiring R]
[Algebra S R]
:
Algebra S (QuadraticAlgebra R a b)
Equations
- One or more equations did not get rendered due to their size.
theorem
QuadraticAlgebra.algebraMap_injective
{R : Type u_1}
{a b : R}
[CommSemiring R]
:
Function.Injective ⇑(algebraMap R (QuadraticAlgebra R a b))
instance
QuadraticAlgebra.instNoZeroSMulDivisors
{R : Type u_1}
{S : Type u_2}
{a b : R}
[CommSemiring R]
[Zero S]
[SMulWithZero S R]
[NoZeroSMulDivisors S R]
:
NoZeroSMulDivisors S (QuadraticAlgebra R a b)
@[simp]
theorem
QuadraticAlgebra.mul_coe_eq_smul
{R : Type u_1}
{a b : R}
[CommSemiring R]
(r : R)
(x : QuadraticAlgebra R a b)
:
@[simp]
instance
QuadraticAlgebra.instCommRing
{R : Type u_1}
{a b : R}
[CommRing R]
:
CommRing (QuadraticAlgebra R a b)
Equations
- One or more equations did not get rendered due to their size.