Documentation

Mathlib.Algebra.QuadraticAlgebra

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 #

Tags #

Quadratic algebra, quadratic extension

structure QuadraticAlgebra (R : Type u) (a b : R) :

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
    theorem QuadraticAlgebra.ext_iff {R : Type u} {a b : R} {x y : QuadraticAlgebra R a b} :
    x = y x.re = y.re x.im = y.im
    theorem QuadraticAlgebra.ext {R : Type u} {a b : R} {x y : QuadraticAlgebra R a b} (re : x.re = y.re) (im : x.im = y.im) :
    x = y
    def QuadraticAlgebra.equivProd {R : Type u_1} (a b : R) :

    The equivalence between quadratic algebra over R and R × R.

    Equations
    Instances For
      @[simp]
      theorem QuadraticAlgebra.equivProd_symm_apply {R : Type u_1} (a b : R) (p : R × R) :
      (equivProd a b).symm p = { re := p.1, im := p.2 }
      @[simp]
      theorem QuadraticAlgebra.mk_eta {R : Type u_1} {a b : R} (z : QuadraticAlgebra R a b) :
      { re := z.re, im := z.im } = z
      def QuadraticAlgebra.coe {R : Type u_1} {a b : R} [Zero R] (x : R) :

      Coercion R → QuadraticAlgebra R a b.

      Equations
      • x = { re := x, im := 0 }
      Instances For
        @[simp]
        theorem QuadraticAlgebra.re_coe {R : Type u_1} {a b : R} (r : R) [Zero R] :
        (↑r).re = r
        @[simp]
        theorem QuadraticAlgebra.im_coe {R : Type u_1} {a b : R} (r : R) [Zero R] :
        (↑r).im = 0
        @[simp]
        theorem QuadraticAlgebra.coe_inj {R : Type u_1} {a b : R} [Zero R] {x y : R} :
        x = y x = y
        instance QuadraticAlgebra.instZero {R : Type u_1} {a b : R} [Zero R] :
        Equations
        @[simp]
        theorem QuadraticAlgebra.re_zero {R : Type u_1} {a b : R} [Zero R] :
        re 0 = 0
        @[simp]
        theorem QuadraticAlgebra.im_zero {R : Type u_1} {a b : R} [Zero R] :
        im 0 = 0
        @[simp]
        theorem QuadraticAlgebra.coe_zero {R : Type u_1} {a b : R} [Zero R] :
        0 = 0
        instance QuadraticAlgebra.instInhabited {R : Type u_1} {a b : R} [Zero R] :
        Equations
        instance QuadraticAlgebra.instOne {R : Type u_1} {a b : R} [Zero R] [One R] :
        Equations
        theorem QuadraticAlgebra.re_one {R : Type u_1} {a b : R} [Zero R] [One R] :
        re 1 = 1
        theorem QuadraticAlgebra.im_one {R : Type u_1} {a b : R} [Zero R] [One R] :
        im 1 = 0
        @[simp]
        theorem QuadraticAlgebra.coe_one {R : Type u_1} {a b : R} [Zero R] [One R] :
        1 = 1
        instance QuadraticAlgebra.instAdd {R : Type u_1} {a b : R} [Add R] :
        Equations
        @[simp]
        theorem QuadraticAlgebra.re_add {R : Type u_1} {a b : R} [Add R] (z w : QuadraticAlgebra R a b) :
        (z + w).re = z.re + w.re
        @[simp]
        theorem QuadraticAlgebra.im_add {R : Type u_1} {a b : R} [Add R] (z w : QuadraticAlgebra R a b) :
        (z + w).im = z.im + w.im
        @[simp]
        theorem QuadraticAlgebra.mk_add_mk {R : Type u_1} {a b : R} [Add R] (z w : QuadraticAlgebra R a b) :
        { re := z.re, im := z.im } + { re := w.re, im := w.im } = { re := z.re + w.re, im := z.im + w.im }
        @[simp]
        theorem QuadraticAlgebra.coe_add {R : Type u_1} {a b : R} [AddZeroClass R] (x y : R) :
        ↑(x + y) = x + y
        instance QuadraticAlgebra.instNeg {R : Type u_1} {a b : R} [Neg R] :
        Equations
        @[simp]
        theorem QuadraticAlgebra.re_neg {R : Type u_1} {a b : R} [Neg R] (z : QuadraticAlgebra R a b) :
        (-z).re = -z.re
        @[simp]
        theorem QuadraticAlgebra.im_neg {R : Type u_1} {a b : R} [Neg R] (z : QuadraticAlgebra R a b) :
        (-z).im = -z.im
        @[simp]
        theorem QuadraticAlgebra.neg_mk {R : Type u_1} {a b : R} [Neg R] (x y : R) :
        -{ re := x, im := y } = { re := -x, im := -y }
        @[simp]
        theorem QuadraticAlgebra.coe_neg {R : Type u_1} {a b : R} [NegZeroClass R] (x : R) :
        ↑(-x) = -x
        instance QuadraticAlgebra.instSub {R : Type u_1} {a b : R} [Sub R] :
        Equations
        @[simp]
        theorem QuadraticAlgebra.re_sub {R : Type u_1} {a b : R} [Sub R] (z w : QuadraticAlgebra R a b) :
        (z - w).re = z.re - w.re
        @[simp]
        theorem QuadraticAlgebra.im_sub {R : Type u_1} {a b : R} [Sub R] (z w : QuadraticAlgebra R a b) :
        (z - w).im = z.im - w.im
        @[simp]
        theorem QuadraticAlgebra.mk_sub_mk {R : Type u_1} {a b : R} [Sub R] (x1 y1 x2 y2 : R) :
        { re := x1, im := y1 } - { re := x2, im := y2 } = { re := x1 - x2, im := y1 - y2 }
        @[simp]
        theorem QuadraticAlgebra.coe_sub {R : Type u_1} {a b : R} (r1 r2 : R) [SubNegZeroMonoid R] :
        ↑(r1 - r2) = r1 - r2
        instance QuadraticAlgebra.instMul {R : Type u_1} {a b : R} [Mul R] [Add R] :
        Equations
        @[simp]
        theorem QuadraticAlgebra.re_mul {R : Type u_1} {a b : R} [Mul R] [Add R] (z w : QuadraticAlgebra R a b) :
        (z * w).re = z.re * w.re + a * z.im * w.im
        @[simp]
        theorem QuadraticAlgebra.im_mul {R : Type u_1} {a b : R} [Mul R] [Add R] (z w : QuadraticAlgebra R a b) :
        (z * w).im = z.re * w.im + z.im * w.re + b * z.im * w.im
        @[simp]
        theorem QuadraticAlgebra.mk_mul_mk {R : Type u_1} {a b : R} [Mul R] [Add R] (x1 y1 x2 y2 : R) :
        { re := x1, im := y1 } * { re := x2, im := y2 } = { re := x1 * x2 + a * y1 * y2, im := x1 * y2 + y1 * x2 + b * y1 * y2 }
        instance QuadraticAlgebra.instSMul {R : Type u_1} {S : Type u_2} {a b : R} [SMul S R] :
        Equations
        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] :
        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] :
        @[simp]
        theorem QuadraticAlgebra.re_smul {R : Type u_1} {S : Type u_2} {a b : R} [SMul S R] (s : S) (z : QuadraticAlgebra R a b) :
        (s z).re = s z.re
        @[simp]
        theorem QuadraticAlgebra.im_smul {R : Type u_1} {S : Type u_2} {a b : R} [SMul S R] (s : S) (z : QuadraticAlgebra R a b) :
        (s z).im = s z.im
        @[simp]
        theorem QuadraticAlgebra.smul_mk {R : Type u_1} {S : Type u_2} {a b : R} [SMul S R] (s : S) (x y : R) :
        s { re := x, im := y } = { re := s x, im := s y }
        instance QuadraticAlgebra.instMulAction {R : Type u_1} {S : Type u_2} {a b : R} [Monoid S] [MulAction S R] :
        Equations
        @[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) :
        ↑(s r) = s r
        instance QuadraticAlgebra.instAddMonoid {R : Type u_1} {a b : R} [AddMonoid R] :
        Equations
        • One or more equations did not get rendered due to their size.
        Equations
        instance QuadraticAlgebra.instModule {R : Type u_1} {S : Type u_2} {a b : R} [Semiring S] [AddCommMonoid R] [Module S R] :
        Equations
        instance QuadraticAlgebra.instAddGroup {R : Type u_1} {a b : R} [AddGroup R] :
        Equations
        • One or more equations did not get rendered due to their size.
        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]
        theorem QuadraticAlgebra.re_natCast {R : Type u_1} {a b : R} [AddCommMonoidWithOne R] (n : ) :
        (↑n).re = n
        @[simp]
        theorem QuadraticAlgebra.im_natCast {R : Type u_1} {a b : R} [AddCommMonoidWithOne R] (n : ) :
        (↑n).im = 0
        theorem QuadraticAlgebra.coe_natCast {R : Type u_1} {a b : R} [AddCommMonoidWithOne R] (n : ) :
        n = n
        theorem QuadraticAlgebra.im_ofNat {R : Type u_1} {a b : R} [AddCommMonoidWithOne R] (n : ) [n.AtLeastTwo] :
        Equations
        • One or more equations did not get rendered due to their size.
        @[simp]
        theorem QuadraticAlgebra.re_intCast {R : Type u_1} {a b : R} [AddCommGroupWithOne R] (n : ) :
        (↑n).re = n
        @[simp]
        theorem QuadraticAlgebra.im_intCast {R : Type u_1} {a b : R} [AddCommGroupWithOne R] (n : ) :
        (↑n).im = 0
        theorem QuadraticAlgebra.coe_intCast {R : Type u_1} {a b : R} [AddCommGroupWithOne R] (n : ) :
        n = n
        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) :
        r * x = r x
        @[simp]
        theorem QuadraticAlgebra.coe_mul {R : Type u_1} {a b : R} [NonUnitalNonAssocSemiring R] (x y : R) :
        ↑(x * y) = x * y
        Equations
        • One or more equations did not get rendered due to their size.
        def QuadraticAlgebra.reₗ {R : Type u_1} (a b : R) [Semiring R] :

        QuadraticAlgebra.re as a LinearMap

        Equations
        Instances For
          @[simp]
          theorem QuadraticAlgebra.reₗ_apply {R : Type u_1} (a b : R) [Semiring R] (self : QuadraticAlgebra R a b) :
          (reₗ a b) self = self.re
          def QuadraticAlgebra.imₗ {R : Type u_1} (a b : R) [Semiring R] :

          QuadraticAlgebra.im as a LinearMap

          Equations
          Instances For
            @[simp]
            theorem QuadraticAlgebra.imₗ_apply {R : Type u_1} (a b : R) [Semiring R] (self : QuadraticAlgebra R a b) :
            (imₗ a b) self = self.im
            def QuadraticAlgebra.linearEquivTuple {R : Type u_1} (a b : R) [Semiring R] :

            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 2R) :
              (linearEquivTuple a b).symm x = { re := x 0, im := x 1 }
              noncomputable def QuadraticAlgebra.basis {R : Type u_1} (a b : R) [Semiring R] :

              QuadraticAlgebra R a b has a basis over R given by 1 and i

              Equations
              Instances For
                @[simp]
                theorem QuadraticAlgebra.basis_repr_apply {R : Type u_1} (a b : R) [Semiring R] (x : QuadraticAlgebra R a b) :
                ((basis a b).repr x) = ![x.re, x.im]
                instance QuadraticAlgebra.instFree {R : Type u_1} (a b : R) [Semiring R] :
                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] :
                Equations
                • One or more equations did not get rendered due to their size.
                theorem QuadraticAlgebra.algebraMap_eq {R : Type u_1} {a b : R} [CommSemiring R] (r : R) :
                (algebraMap R (QuadraticAlgebra R a b)) r = { re := r, im := 0 }
                @[simp]
                theorem QuadraticAlgebra.coe_pow {R : Type u_1} {a b : R} [CommSemiring R] (n : ) (r : R) :
                ↑(r ^ n) = r ^ n
                theorem QuadraticAlgebra.mul_coe_eq_smul {R : Type u_1} {a b : R} [CommSemiring R] (r : R) (x : QuadraticAlgebra R a b) :
                x * r = r x
                @[simp]
                theorem QuadraticAlgebra.coe_algebraMap {R : Type u_1} {a b : R} [CommSemiring R] :
                theorem QuadraticAlgebra.smul_coe {R : Type u_1} {a b : R} [CommSemiring R] (r1 r2 : R) :
                r1 r2 = ↑(r1 * r2)
                instance QuadraticAlgebra.instCommRing {R : Type u_1} {a b : R} [CommRing R] :
                Equations
                • One or more equations did not get rendered due to their size.