Documentation

Mathlib.LinearAlgebra.Matrix.GeneralLinearGroup.FinTwo

The group GL (Fin 2) R #

def Matrix.IsParabolic {R : Type u_1} [CommRing R] (m : Matrix (Fin 2) (Fin 2) R) :

A 2 × 2 matrix is parabolic if it is non-scalar and its discriminant is 0.

Equations
Instances For
    @[simp]
    theorem Matrix.disc_conj {R : Type u_1} [CommRing R] [Nontrivial R] {m : Matrix (Fin 2) (Fin 2) R} (g : GL (Fin 2) R) :
    (g * m * (↑g)⁻¹).disc = m.disc
    @[simp]
    theorem Matrix.disc_conj' {R : Type u_1} [CommRing R] [Nontrivial R] {m : Matrix (Fin 2) (Fin 2) R} (g : GL (Fin 2) R) :
    ((↑g)⁻¹ * m * g).disc = m.disc
    @[simp]
    theorem Matrix.isParabolic_conj_iff {R : Type u_1} [CommRing R] [Nontrivial R] {m : Matrix (Fin 2) (Fin 2) R} (g : GL (Fin 2) R) :
    (g * m * (↑g)⁻¹).IsParabolic m.IsParabolic
    @[simp]
    theorem Matrix.isParabolic_conj'_iff {R : Type u_1} [CommRing R] [Nontrivial R] {m : Matrix (Fin 2) (Fin 2) R} (g : GL (Fin 2) R) :
    ((↑g)⁻¹ * m * g).IsParabolic m.IsParabolic
    theorem Matrix.isParabolic_iff_of_upperTriangular {R : Type u_1} [CommRing R] [Nontrivial R] {m : Matrix (Fin 2) (Fin 2) R} [IsReduced R] (hm : m 1 0 = 0) :
    m.IsParabolic m 0 0 = m 1 1 m 0 1 0
    theorem Matrix.sub_scalar_sq_eq_disc {K : Type u_1} [Field K] {m : Matrix (Fin 2) (Fin 2) K} [NeZero 2] :
    (m - (scalar (Fin 2)) (m.trace / 2)) ^ 2 = (scalar (Fin 2)) (m.disc / 4)
    def Matrix.parabolicEigenvalue {K : Type u_1} [Field K] (m : Matrix (Fin 2) (Fin 2) K) :
    K

    The unique eigenvalue of a parabolic matrix (junk if m is not parabolic).

    Equations
    Instances For
      theorem Matrix.IsParabolic.sub_eigenvalue_sq_eq_zero {K : Type u_1} [Field K] {m : Matrix (Fin 2) (Fin 2) K} [NeZero 2] (hm : m.IsParabolic) :
      (m - (scalar (Fin 2)) m.parabolicEigenvalue) ^ 2 = 0
      theorem Matrix.isParabolic_iff_exists {K : Type u_1} [Field K] {m : Matrix (Fin 2) (Fin 2) K} [NeZero 2] :
      m.IsParabolic ∃ (a : K) (n : Matrix (Fin 2) (Fin 2) K), m = (scalar (Fin 2)) a + n n 0 n ^ 2 = 0

      Characterization of parabolic elements: they have the form a + m where a is scalar and m is nonzero and nilpotent.

      def Matrix.IsHyperbolic {R : Type u_1} [CommRing R] [Preorder R] (m : Matrix (Fin 2) (Fin 2) R) :

      A 2 × 2 matrix is hyperbolic if its discriminant is strictly positive.

      Equations
      Instances For
        def Matrix.IsElliptic {R : Type u_1} [CommRing R] [Preorder R] (m : Matrix (Fin 2) (Fin 2) R) :

        A 2 × 2 matrix is elliptic if its discriminant is strictly negative.

        Equations
        Instances For
          theorem Matrix.isHyperbolic_conj_iff {R : Type u_1} [CommRing R] [Nontrivial R] [Preorder R] {m : Matrix (Fin 2) (Fin 2) R} (g : GL (Fin 2) R) :
          theorem Matrix.isHyperbolic_conj'_iff {R : Type u_1} [CommRing R] [Nontrivial R] [Preorder R] {m : Matrix (Fin 2) (Fin 2) R} (g : GL (Fin 2) R) :
          theorem Matrix.isElliptic_conj_iff {R : Type u_1} [CommRing R] [Nontrivial R] [Preorder R] {m : Matrix (Fin 2) (Fin 2) R} (g : GL (Fin 2) R) :
          (g * m * (↑g)⁻¹).IsElliptic m.IsElliptic
          theorem Matrix.isElliptic_conj'_iff {R : Type u_1} [CommRing R] [Nontrivial R] [Preorder R] {m : Matrix (Fin 2) (Fin 2) R} (g : GL (Fin 2) R) :
          ((↑g)⁻¹ * m * g).IsElliptic m.IsElliptic

          The map sending x to [1, x; 0, 1] (bundled as an AddChar).

          Equations
          Instances For
            @[simp]
            theorem Matrix.GeneralLinearGroup.upperRightHom_apply {R : Type u_1} [Ring R] (x : R) :
            upperRightHom x = { val := !![1, x; 0, 1], inv := !![1, -x; 0, 1], val_inv := , inv_val := }
            @[reducible, inline]
            abbrev Matrix.GeneralLinearGroup.IsParabolic {R : Type u_1} [CommRing R] (g : GL (Fin 2) R) :

            Synonym of Matrix.IsParabolic, for dot-notation.

            Equations
            Instances For
              @[reducible, inline]
              abbrev Matrix.GeneralLinearGroup.IsElliptic {R : Type u_1} [CommRing R] [Preorder R] (g : GL (Fin 2) R) :

              Synonym of Matrix.IsElliptic, for dot-notation.

              Equations
              Instances For
                @[reducible, inline]

                Synonym of Matrix.IsHyperbolic, for dot-notation.

                Equations
                Instances For
                  noncomputable def Matrix.GeneralLinearGroup.fixpointPolynomial {R : Type u_1} [CommRing R] (g : GL (Fin 2) R) :

                  Polynomial whose roots are the fixed points of g considered as a Möbius transformation.

                  See Matrix.GeneralLinearGroup.fixpointPolynomial_aeval_eq_zero_iff.

                  Equations
                  Instances For

                    The fixed-point polynomial is identically zero iff g is scalar.

                    theorem Matrix.GeneralLinearGroup.IsParabolic.pow {K : Type u_2} [Field K] {g : GL (Fin 2) K} (hg : g.IsParabolic) [CharZero K] {n : } (hn : n 0) :

                    A non-zero power of a parabolic element is parabolic.

                    theorem Matrix.GeneralLinearGroup.isParabolic_iff_of_upperTriangular {K : Type u_2} [Field K] {g : GL (Fin 2) K} (hg : g 1 0 = 0) :
                    g.IsParabolic g 0 0 = g 1 1 g 0 1 0
                    theorem Matrix.GeneralLinearGroup.isParabolic_iff_of_upperTriangular_of_det {K : Type u_2} [Field K] [LinearOrder K] [IsStrictOrderedRing K] {g : GL (Fin 2) K} (h_det : det g = 1 det g = -1) (hg10 : g 1 0 = 0) :
                    g.IsParabolic (∃ (x : K), x 0 g = upperRightHom x) ∃ (x : K), x 0 g = -upperRightHom x

                    Specialized version of isParabolic_iff_of_upperTriangular intended for use with discrete subgroups of GL(2, ℝ).