Documentation

Mathlib.Algebra.Hom.Equiv.Units.Basic

Multiplicative and additive equivalence acting on units. #

theorem toAddUnits.proof_1 {G : Type u_1} [AddGroup G] :
∀ (x : G), (fun x => x) ((fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) x) = (fun x => x) ((fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) x)
theorem toAddUnits.proof_3 {G : Type u_1} [AddGroup G] :
∀ (x x_1 : G), Equiv.toFun { toFun := fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }, invFun := fun x => x, left_inv := (_ : ∀ (x : G), (fun x => x) ((fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) x) = (fun x => x) ((fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) x)), right_inv := (_ : ∀ (x : AddUnits G), (fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) ((fun x => x) x) = x) } (x + x_1) = Equiv.toFun { toFun := fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }, invFun := fun x => x, left_inv := (_ : ∀ (x : G), (fun x => x) ((fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) x) = (fun x => x) ((fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) x)), right_inv := (_ : ∀ (x : AddUnits G), (fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) ((fun x => x) x) = x) } x + Equiv.toFun { toFun := fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }, invFun := fun x => x, left_inv := (_ : ∀ (x : G), (fun x => x) ((fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) x) = (fun x => x) ((fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) x)), right_inv := (_ : ∀ (x : AddUnits G), (fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) ((fun x => x) x) = x) } x_1
theorem toAddUnits.proof_2 {G : Type u_1} [AddGroup G] :
∀ (x : AddUnits G), (fun x => { val := x, neg := -x, val_neg := (_ : x + -x = 0), neg_val := (_ : -x + x = 0) }) ((fun x => x) x) = x
def toAddUnits {G : Type u_10} [AddGroup G] :

An additive group is isomorphic to its group of additive units

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    def toUnits {G : Type u_10} [Group G] :
    G ≃* Gˣ

    A group is isomorphic to its group of units.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem coe_toAddUnits {G : Type u_10} [AddGroup G] (g : G) :
      ↑(toAddUnits g) = g
      @[simp]
      theorem coe_toUnits {G : Type u_10} [Group G] (g : G) :
      ↑(toUnits g) = g
      def Units.mapEquiv {M : Type u_6} {N : Type u_7} [Monoid M] [Monoid N] (h : M ≃* N) :

      A multiplicative equivalence of monoids defines a multiplicative equivalence of their groups of units.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        @[simp]
        theorem Units.coe_mapEquiv {M : Type u_6} {N : Type u_7} [Monoid M] [Monoid N] (h : M ≃* N) (x : Mˣ) :
        ↑(↑(Units.mapEquiv h) x) = h x
        def AddUnits.addLeft {M : Type u_6} [AddMonoid M] (u : AddUnits M) :

        Left addition of an additive unit is a permutation of the underlying type.

        Equations
        • AddUnits.addLeft u = { toFun := fun x => u + x, invFun := fun x => ↑(-u) + x, left_inv := (_ : ∀ (b : M), ↑(-u) + (u + b) = b), right_inv := (_ : ∀ (b : M), u + (↑(-u) + b) = b) }
        Instances For
          @[simp]
          theorem AddUnits.addLeft_apply {M : Type u_6} [AddMonoid M] (u : AddUnits M) :
          ↑(AddUnits.addLeft u) = fun x => u + x
          @[simp]
          theorem Units.mulLeft_apply {M : Type u_6} [Monoid M] (u : Mˣ) :
          ↑(Units.mulLeft u) = fun x => u * x
          def Units.mulLeft {M : Type u_6} [Monoid M] (u : Mˣ) :

          Left multiplication by a unit of a monoid is a permutation of the underlying type.

          Equations
          • Units.mulLeft u = { toFun := fun x => u * x, invFun := fun x => u⁻¹ * x, left_inv := (_ : ∀ (b : M), u⁻¹ * (u * b) = b), right_inv := (_ : ∀ (b : M), u * (u⁻¹ * b) = b) }
          Instances For
            @[simp]
            @[simp]
            theorem Units.mulLeft_symm {M : Type u_6} [Monoid M] (u : Mˣ) :
            theorem AddUnits.addLeft_bijective {M : Type u_6} [AddMonoid M] (a : AddUnits M) :
            Function.Bijective fun x => a + x
            theorem Units.mulLeft_bijective {M : Type u_6} [Monoid M] (a : Mˣ) :
            Function.Bijective fun x => a * x
            def AddUnits.addRight {M : Type u_6} [AddMonoid M] (u : AddUnits M) :

            Right addition of an additive unit is a permutation of the underlying type.

            Equations
            • AddUnits.addRight u = { toFun := fun x => x + u, invFun := fun x => x + ↑(-u), left_inv := (_ : ∀ (x : M), x + u + ↑(-u) = x), right_inv := (_ : ∀ (x : M), x + ↑(-u) + u = x) }
            Instances For
              @[simp]
              theorem AddUnits.addRight_apply {M : Type u_6} [AddMonoid M] (u : AddUnits M) :
              ↑(AddUnits.addRight u) = fun x => x + u
              @[simp]
              theorem Units.mulRight_apply {M : Type u_6} [Monoid M] (u : Mˣ) :
              ↑(Units.mulRight u) = fun x => x * u
              def Units.mulRight {M : Type u_6} [Monoid M] (u : Mˣ) :

              Right multiplication by a unit of a monoid is a permutation of the underlying type.

              Equations
              • Units.mulRight u = { toFun := fun x => x * u, invFun := fun x => x * u⁻¹, left_inv := (_ : ∀ (x : M), x * u * u⁻¹ = x), right_inv := (_ : ∀ (x : M), x * u⁻¹ * u = x) }
              Instances For
                @[simp]
                @[simp]
                theorem Units.mulRight_symm {M : Type u_6} [Monoid M] (u : Mˣ) :
                theorem AddUnits.addRight_bijective {M : Type u_6} [AddMonoid M] (a : AddUnits M) :
                Function.Bijective fun x => x + a
                theorem Units.mulRight_bijective {M : Type u_6} [Monoid M] (a : Mˣ) :
                Function.Bijective fun x => x * a
                def Equiv.addLeft {G : Type u_10} [AddGroup G] (a : G) :

                Left addition in an AddGroup is a permutation of the underlying type.

                Equations
                Instances For
                  def Equiv.mulLeft {G : Type u_10} [Group G] (a : G) :

                  Left multiplication in a Group is a permutation of the underlying type.

                  Equations
                  Instances For
                    @[simp]
                    theorem Equiv.coe_addLeft {G : Type u_10} [AddGroup G] (a : G) :
                    ↑(Equiv.addLeft a) = (fun x x_1 => x + x_1) a
                    @[simp]
                    theorem Equiv.coe_mulLeft {G : Type u_10} [Group G] (a : G) :
                    ↑(Equiv.mulLeft a) = (fun x x_1 => x * x_1) a
                    theorem Equiv.addLeft_symm_apply {G : Type u_10} [AddGroup G] (a : G) :
                    (Equiv.addLeft a).symm = fun x => -a + x

                    Extra simp lemma that dsimp can use. simp will never use this.

                    @[simp]
                    theorem Equiv.mulLeft_symm_apply {G : Type u_10} [Group G] (a : G) :
                    (Equiv.mulLeft a).symm = fun x => a⁻¹ * x

                    Extra simp lemma that dsimp can use. simp will never use this.

                    @[simp]
                    theorem Equiv.addLeft_symm {G : Type u_10} [AddGroup G] (a : G) :
                    @[simp]
                    theorem Equiv.mulLeft_symm {G : Type u_10} [Group G] (a : G) :
                    theorem AddGroup.addLeft_bijective {G : Type u_10} [AddGroup G] (a : G) :
                    Function.Bijective fun x => a + x
                    theorem Group.mulLeft_bijective {G : Type u_10} [Group G] (a : G) :
                    Function.Bijective fun x => a * x
                    def Equiv.addRight {G : Type u_10} [AddGroup G] (a : G) :

                    Right addition in an AddGroup is a permutation of the underlying type.

                    Equations
                    Instances For
                      def Equiv.mulRight {G : Type u_10} [Group G] (a : G) :

                      Right multiplication in a Group is a permutation of the underlying type.

                      Equations
                      Instances For
                        @[simp]
                        theorem Equiv.coe_addRight {G : Type u_10} [AddGroup G] (a : G) :
                        ↑(Equiv.addRight a) = fun x => x + a
                        @[simp]
                        theorem Equiv.coe_mulRight {G : Type u_10} [Group G] (a : G) :
                        ↑(Equiv.mulRight a) = fun x => x * a
                        @[simp]
                        theorem Equiv.addRight_symm {G : Type u_10} [AddGroup G] (a : G) :
                        @[simp]
                        theorem Equiv.mulRight_symm {G : Type u_10} [Group G] (a : G) :
                        theorem Equiv.addRight_symm_apply {G : Type u_10} [AddGroup G] (a : G) :
                        (Equiv.addRight a).symm = fun x => x + -a

                        Extra simp lemma that dsimp can use. simp will never use this.

                        @[simp]
                        theorem Equiv.mulRight_symm_apply {G : Type u_10} [Group G] (a : G) :
                        (Equiv.mulRight a).symm = fun x => x * a⁻¹

                        Extra simp lemma that dsimp can use. simp will never use this.

                        theorem AddGroup.addRight_bijective {G : Type u_10} [AddGroup G] (a : G) :
                        Function.Bijective fun x => x + a
                        theorem Group.mulRight_bijective {G : Type u_10} [Group G] (a : G) :
                        Function.Bijective fun x => x * a
                        def Equiv.subLeft {G : Type u_10} [AddGroup G] (a : G) :
                        G G

                        A version of Equiv.addLeft a (-b) that is defeq to a - b.

                        Equations
                        • Equiv.subLeft a = { toFun := fun b => a - b, invFun := fun b => -b + a, left_inv := (_ : ∀ (b : G), -(a - b) + a = b), right_inv := (_ : ∀ (b : G), a - (-b + a) = b) }
                        Instances For
                          theorem Equiv.subLeft.proof_1 {G : Type u_1} [AddGroup G] (a : G) (b : G) :
                          -(a - b) + a = b
                          theorem Equiv.subLeft.proof_2 {G : Type u_1} [AddGroup G] (a : G) (b : G) :
                          a - (-b + a) = b
                          @[simp]
                          theorem Equiv.divLeft_apply {G : Type u_10} [Group G] (a : G) (b : G) :
                          ↑(Equiv.divLeft a) b = a / b
                          @[simp]
                          theorem Equiv.divLeft_symm_apply {G : Type u_10} [Group G] (a : G) (b : G) :
                          (Equiv.divLeft a).symm b = b⁻¹ * a
                          @[simp]
                          theorem Equiv.subLeft_symm_apply {G : Type u_10} [AddGroup G] (a : G) (b : G) :
                          (Equiv.subLeft a).symm b = -b + a
                          @[simp]
                          theorem Equiv.subLeft_apply {G : Type u_10} [AddGroup G] (a : G) (b : G) :
                          ↑(Equiv.subLeft a) b = a - b
                          def Equiv.divLeft {G : Type u_10} [Group G] (a : G) :
                          G G

                          A version of Equiv.mulLeft a b⁻¹ that is defeq to a / b.

                          Equations
                          Instances For
                            def Equiv.subRight {G : Type u_10} [AddGroup G] (a : G) :
                            G G

                            A version of Equiv.addRight (-a) b that is defeq to b - a.

                            Equations
                            • Equiv.subRight a = { toFun := fun b => b - a, invFun := fun b => b + a, left_inv := (_ : ∀ (b : G), b - a + a = b), right_inv := (_ : ∀ (b : G), b + a - a = b) }
                            Instances For
                              theorem Equiv.subRight.proof_2 {G : Type u_1} [AddGroup G] (a : G) (b : G) :
                              b + a - a = b
                              theorem Equiv.subRight.proof_1 {G : Type u_1} [AddGroup G] (a : G) (b : G) :
                              b - a + a = b
                              @[simp]
                              theorem Equiv.divRight_symm_apply {G : Type u_10} [Group G] (a : G) (b : G) :
                              (Equiv.divRight a).symm b = b * a
                              @[simp]
                              theorem Equiv.subRight_apply {G : Type u_10} [AddGroup G] (a : G) (b : G) :
                              ↑(Equiv.subRight a) b = b - a
                              @[simp]
                              theorem Equiv.subRight_symm_apply {G : Type u_10} [AddGroup G] (a : G) (b : G) :
                              (Equiv.subRight a).symm b = b + a
                              @[simp]
                              theorem Equiv.divRight_apply {G : Type u_10} [Group G] (a : G) (b : G) :
                              ↑(Equiv.divRight a) b = b / a
                              def Equiv.divRight {G : Type u_10} [Group G] (a : G) :
                              G G

                              A version of Equiv.mulRight a⁻¹ b that is defeq to b / a.

                              Equations
                              • Equiv.divRight a = { toFun := fun b => b / a, invFun := fun b => b * a, left_inv := (_ : ∀ (b : G), b / a * a = b), right_inv := (_ : ∀ (b : G), b * a / a = b) }
                              Instances For

                                When the AddGroup is commutative, Equiv.neg is an AddEquiv.

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For
                                  @[simp]
                                  theorem AddEquiv.neg_apply (G : Type u_12) [SubtractionCommMonoid G] :
                                  ∀ (a : G), ↑(AddEquiv.neg G) a = -a
                                  @[simp]
                                  theorem MulEquiv.inv_apply (G : Type u_12) [DivisionCommMonoid G] :
                                  ∀ (a : G), ↑(MulEquiv.inv G) a = a⁻¹
                                  def MulEquiv.inv (G : Type u_12) [DivisionCommMonoid G] :
                                  G ≃* G

                                  In a DivisionCommMonoid, Equiv.inv is a MulEquiv. There is a variant of this MulEquiv.inv' G : G ≃* Gᵐᵒᵖ for the non-commutative case.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For