Documentation

Mathlib.LinearAlgebra.Orientation

Orientations of modules #

This file defines orientations of modules.

Main definitions #

Implementation notes #

Orientation is defined for an arbitrary index type, but the main intended use case is when that index type is a Fintype and there exists a basis of the same cardinality.

References #

@[inline, reducible]
abbrev Orientation (R : Type u_1) [StrictOrderedCommSemiring R] (M : Type u_2) [AddCommMonoid M] [Module R M] (ι : Type u_4) :
Type (max (max u_2 u_1) u_4)

An orientation of a module, intended to be used when ι is a Fintype with the same cardinality as a basis.

Equations
Instances For
    class Module.Oriented (R : Type u_1) [StrictOrderedCommSemiring R] (M : Type u_2) [AddCommMonoid M] [Module R M] (ι : Type u_4) :
    Type (max (max u_1 u_2) u_4)

    A type class fixing an orientation of a module.

    • positiveOrientation : Orientation R M ι

      Fix a positive orientation.

    Instances
      def Orientation.map {R : Type u_1} [StrictOrderedCommSemiring R] {M : Type u_2} [AddCommMonoid M] [Module R M] {N : Type u_3} [AddCommMonoid N] [Module R N] (ι : Type u_4) (e : M ≃ₗ[R] N) :

      An equivalence between modules implies an equivalence between orientations.

      Equations
      Instances For
        @[simp]
        theorem Orientation.map_apply {R : Type u_1} [StrictOrderedCommSemiring R] {M : Type u_2} [AddCommMonoid M] [Module R M] {N : Type u_3} [AddCommMonoid N] [Module R N] (ι : Type u_4) (e : M ≃ₗ[R] N) (v : M [⋀^ι]→ₗ[R] R) (hv : v 0) :
        @[simp]
        @[simp]
        theorem Orientation.map_symm {R : Type u_1} [StrictOrderedCommSemiring R] {M : Type u_2} [AddCommMonoid M] [Module R M] {N : Type u_3} [AddCommMonoid N] [Module R N] (ι : Type u_4) (e : M ≃ₗ[R] N) :
        def Orientation.reindex (R : Type u_1) [StrictOrderedCommSemiring R] (M : Type u_2) [AddCommMonoid M] [Module R M] {ι : Type u_4} {ι' : Type u_5} (e : ι ι') :
        Orientation R M ι Orientation R M ι'

        An equivalence between indices implies an equivalence between orientations.

        Equations
        Instances For
          @[simp]
          theorem Orientation.reindex_apply (R : Type u_1) [StrictOrderedCommSemiring R] (M : Type u_2) [AddCommMonoid M] [Module R M] {ι : Type u_4} {ι' : Type u_5} (e : ι ι') (v : M [⋀^ι]→ₗ[R] R) (hv : v 0) :
          @[simp]
          theorem Orientation.reindex_symm (R : Type u_1) [StrictOrderedCommSemiring R] (M : Type u_2) [AddCommMonoid M] [Module R M] {ι : Type u_4} {ι' : Type u_5} (e : ι ι') :
          (Orientation.reindex R M e).symm = Orientation.reindex R M e.symm
          instance IsEmpty.oriented {R : Type u_1} [StrictOrderedCommSemiring R] {M : Type u_2} [AddCommMonoid M] [Module R M] (ι : Type u_4) [IsEmpty ι] :

          A module is canonically oriented with respect to an empty index type.

          Equations
          @[simp]
          theorem Orientation.map_positiveOrientation_of_isEmpty {R : Type u_1} [StrictOrderedCommSemiring R] {M : Type u_2} [AddCommMonoid M] [Module R M] {N : Type u_3} [AddCommMonoid N] [Module R N] (ι : Type u_4) [IsEmpty ι] (f : M ≃ₗ[R] N) :
          (Orientation.map ι f) positiveOrientation = positiveOrientation
          @[simp]
          theorem Orientation.map_of_isEmpty {R : Type u_1} [StrictOrderedCommSemiring R] {M : Type u_2} [AddCommMonoid M] [Module R M] (ι : Type u_4) [IsEmpty ι] (x : Orientation R M ι) (f : M ≃ₗ[R] M) :
          (Orientation.map ι f) x = x
          @[simp]
          theorem Orientation.map_neg {R : Type u_1} [StrictOrderedCommRing R] {M : Type u_2} {N : Type u_3} [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] {ι : Type u_4} (f : M ≃ₗ[R] N) (x : Orientation R M ι) :
          (Orientation.map ι f) (-x) = -(Orientation.map ι f) x
          @[simp]
          theorem Orientation.reindex_neg {R : Type u_1} [StrictOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_4} {ι' : Type u_5} (e : ι ι') (x : Orientation R M ι) :
          theorem Basis.map_orientation_eq_det_inv_smul {R : Type u_1} [StrictOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_4} [Finite ι] (e : Basis ι R M) (x : Orientation R M ι) (f : M ≃ₗ[R] M) :
          (Orientation.map ι f) x = (LinearEquiv.det f)⁻¹ x

          The value of Orientation.map when the index type has the cardinality of a basis, in terms of f.det.

          def Basis.orientation {R : Type u_1} [StrictOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_4} [Fintype ι] [DecidableEq ι] (e : Basis ι R M) :

          The orientation given by a basis.

          Equations
          Instances For
            theorem Basis.orientation_map {R : Type u_1} [StrictOrderedCommRing R] {M : Type u_2} {N : Type u_3} [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] {ι : Type u_4} [Fintype ι] [DecidableEq ι] (e : Basis ι R M) (f : M ≃ₗ[R] N) :
            theorem Basis.orientation_reindex {R : Type u_1} [StrictOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_4} {ι' : Type u_5} [Fintype ι] [DecidableEq ι] [Fintype ι'] [DecidableEq ι'] (e : Basis ι R M) (eι : ι ι') :
            theorem Basis.orientation_unitsSMul {R : Type u_1} [StrictOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_4} [Fintype ι] [DecidableEq ι] (e : Basis ι R M) (w : ιRˣ) :
            Basis.orientation (Basis.unitsSMul e w) = (Finset.prod Finset.univ fun (i : ι) => w i)⁻¹ Basis.orientation e

            The orientation given by a basis derived using units_smul, in terms of the product of those units.

            @[simp]
            theorem Basis.orientation_isEmpty {R : Type u_1} [StrictOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_4} [Fintype ι] [DecidableEq ι] [IsEmpty ι] (b : Basis ι R M) :
            Basis.orientation b = positiveOrientation
            theorem Orientation.eq_or_eq_neg_of_isEmpty {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [IsEmpty ι] (o : Orientation R M ι) :
            o = positiveOrientation o = -positiveOrientation

            A module M over a linearly ordered commutative ring has precisely two "orientations" with respect to an empty index type. (Note that these are only orientations of M of in the conventional mathematical sense if M is zero-dimensional.)

            theorem Basis.orientation_eq_iff_det_pos {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] (e₁ : Basis ι R M) (e₂ : Basis ι R M) :
            Basis.orientation e₁ = Basis.orientation e₂ 0 < (Basis.det e₁) e₂

            The orientations given by two bases are equal if and only if the determinant of one basis with respect to the other is positive.

            theorem Basis.orientation_eq_or_eq_neg {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] (e : Basis ι R M) (x : Orientation R M ι) :

            Given a basis, any orientation equals the orientation given by that basis or its negation.

            theorem Basis.orientation_ne_iff_eq_neg {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] (e : Basis ι R M) (x : Orientation R M ι) :

            Given a basis, an orientation equals the negation of that given by that basis if and only if it does not equal that given by that basis.

            theorem Basis.orientation_comp_linearEquiv_eq_iff_det_pos {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] (e : Basis ι R M) (f : M ≃ₗ[R] M) :
            Basis.orientation (Basis.map e f) = Basis.orientation e 0 < LinearMap.det f

            Composing a basis with a linear equiv gives the same orientation if and only if the determinant is positive.

            theorem Basis.orientation_comp_linearEquiv_eq_neg_iff_det_neg {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] (e : Basis ι R M) (f : M ≃ₗ[R] M) :
            Basis.orientation (Basis.map e f) = -Basis.orientation e LinearMap.det f < 0

            Composing a basis with a linear equiv gives the negation of that orientation if and only if the determinant is negative.

            @[simp]
            theorem Basis.orientation_neg_single {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] (e : Basis ι R M) (i : ι) :

            Negating a single basis vector (represented using units_smul) negates the corresponding orientation.

            def Basis.adjustToOrientation {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) :
            Basis ι R M

            Given a basis and an orientation, return a basis giving that orientation: either the original basis, or one constructed by negating a single (arbitrary) basis vector.

            Equations
            Instances For
              @[simp]
              theorem Basis.orientation_adjustToOrientation {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) :

              adjust_to_orientation gives a basis with the required orientation.

              theorem Basis.adjustToOrientation_apply_eq_or_eq_neg {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) (i : ι) :

              Every basis vector from adjust_to_orientation is either that from the original basis or its negation.

              @[simp]
              theorem Basis.abs_det_adjustToOrientation {R : Type u_1} [LinearOrderedCommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [DecidableEq ι] [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) (v : ιM) :
              theorem Orientation.eq_or_eq_neg {R : Type u_1} [LinearOrderedField R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [_i : FiniteDimensional R M] (x₁ : Orientation R M ι) (x₂ : Orientation R M ι) (h : Fintype.card ι = FiniteDimensional.finrank R M) :
              x₁ = x₂ x₁ = -x₂

              If the index type has cardinality equal to the finite dimension, any two orientations are equal or negations.

              theorem Orientation.ne_iff_eq_neg {R : Type u_1} [LinearOrderedField R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [_i : FiniteDimensional R M] (x₁ : Orientation R M ι) (x₂ : Orientation R M ι) (h : Fintype.card ι = FiniteDimensional.finrank R M) :
              x₁ x₂ x₁ = -x₂

              If the index type has cardinality equal to the finite dimension, an orientation equals the negation of another orientation if and only if they are not equal.

              theorem Orientation.map_eq_det_inv_smul {R : Type u_1} [LinearOrderedField R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [_i : FiniteDimensional R M] (x : Orientation R M ι) (f : M ≃ₗ[R] M) (h : Fintype.card ι = FiniteDimensional.finrank R M) :
              (Orientation.map ι f) x = (LinearEquiv.det f)⁻¹ x

              The value of Orientation.map when the index type has cardinality equal to the finite dimension, in terms of f.det.

              theorem Orientation.map_eq_iff_det_pos {R : Type u_1} [LinearOrderedField R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [_i : FiniteDimensional R M] (x : Orientation R M ι) (f : M ≃ₗ[R] M) (h : Fintype.card ι = FiniteDimensional.finrank R M) :
              (Orientation.map ι f) x = x 0 < LinearMap.det f

              If the index type has cardinality equal to the finite dimension, composing an alternating map with the same linear equiv on each argument gives the same orientation if and only if the determinant is positive.

              theorem Orientation.map_eq_neg_iff_det_neg {R : Type u_1} [LinearOrderedField R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] (x : Orientation R M ι) (f : M ≃ₗ[R] M) (h : Fintype.card ι = FiniteDimensional.finrank R M) :
              (Orientation.map ι f) x = -x LinearMap.det f < 0

              If the index type has cardinality equal to the finite dimension, composing an alternating map with the same linear equiv on each argument gives the negation of that orientation if and only if the determinant is negative.

              def Orientation.someBasis {R : Type u_1} [LinearOrderedField R] {M : Type u_2} [AddCommGroup M] [Module R M] {ι : Type u_3} [Fintype ι] [_i : FiniteDimensional R M] [Nonempty ι] [DecidableEq ι] (x : Orientation R M ι) (h : Fintype.card ι = FiniteDimensional.finrank R M) :
              Basis ι R M

              If the index type has cardinality equal to the finite dimension, a basis with the given orientation.

              Equations
              Instances For
                @[simp]

                some_basis gives a basis with the required orientation.