Documentation

Mathlib.Topology.Algebra.Module.Equiv

Continuous linear equivalences #

Continuous semilinear / linear / star-linear equivalences between topological modules are denoted by M ≃SL[σ] M₂, M ≃L[R] M₂ and M ≃L⋆[R] M₂.

structure ContinuousLinearEquiv {R : Type u_3} {S : Type u_4} [Semiring R] [Semiring S] (σ : R →+* S) {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : Type u_5) [TopologicalSpace M] [AddCommMonoid M] (M₂ : Type u_6) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] [Module S M₂] extends M ≃ₛₗ[σ] M₂ :
Type (max u_5 u_6)

Continuous linear equivalences between modules. We only put the type classes that are necessary for the definition, although in applications M and M₂ will be topological modules over the topological semiring R.

  • toFun : MM₂
  • map_add' (x y : M) : (↑self.toLinearEquiv).toFun (x + y) = (↑self.toLinearEquiv).toFun x + (↑self.toLinearEquiv).toFun y
  • map_smul' (m : R) (x : M) : (↑self.toLinearEquiv).toFun (m x) = σ m (↑self.toLinearEquiv).toFun x
  • invFun : M₂M
  • left_inv : Function.LeftInverse self.invFun (↑self.toLinearEquiv).toFun
  • right_inv : Function.RightInverse self.invFun (↑self.toLinearEquiv).toFun
  • continuous_toFun : Continuous (↑self.toLinearEquiv).toFun

    Continuous linear equivalences between modules. We only put the type classes that are necessary for the definition, although in applications M and M₂ will be topological modules over the topological semiring R.

  • continuous_invFun : Continuous self.invFun

    Continuous linear equivalences between modules. We only put the type classes that are necessary for the definition, although in applications M and M₂ will be topological modules over the topological semiring R.

Instances For

    Continuous linear equivalences between modules. We only put the type classes that are necessary for the definition, although in applications M and M₂ will be topological modules over the topological semiring R.

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

      Continuous linear equivalences between modules. We only put the type classes that are necessary for the definition, although in applications M and M₂ will be topological modules over the topological semiring R.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        class ContinuousSemilinearEquivClass (F : Type u_3) {R : outParam (Type u_4)} {S : outParam (Type u_5)} [Semiring R] [Semiring S] (σ : outParam (R →+* S)) {σ' : outParam (S →+* R)} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : outParam (Type u_6)) [TopologicalSpace M] [AddCommMonoid M] (M₂ : outParam (Type u_7)) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] [Module S M₂] [EquivLike F M M₂] extends SemilinearEquivClass F σ M M₂ :

        ContinuousSemilinearEquivClass F σ M M₂ asserts F is a type of bundled continuous σ-semilinear equivs M → M₂. See also ContinuousLinearEquivClass F R M M₂ for the case where σ is the identity map on R. A map f between an R-module and an S-module over a ring homomorphism σ : R →+* S is semilinear if it satisfies the two properties f (x + y) = f x + f y and f (c • x) = (σ c) • f x.

        • map_add (f : F) (a b : M) : f (a + b) = f a + f b
        • map_smulₛₗ (f : F) (r : R) (x : M) : f (r x) = σ r f x
        • map_continuous (f : F) : Continuous f

          ContinuousSemilinearEquivClass F σ M M₂ asserts F is a type of bundled continuous σ-semilinear equivs M → M₂. See also ContinuousLinearEquivClass F R M M₂ for the case where σ is the identity map on R. A map f between an R-module and an S-module over a ring homomorphism σ : R →+* S is semilinear if it satisfies the two properties f (x + y) = f x + f y and f (c • x) = (σ c) • f x.

        • inv_continuous (f : F) : Continuous (EquivLike.inv f)

          ContinuousSemilinearEquivClass F σ M M₂ asserts F is a type of bundled continuous σ-semilinear equivs M → M₂. See also ContinuousLinearEquivClass F R M M₂ for the case where σ is the identity map on R. A map f between an R-module and an S-module over a ring homomorphism σ : R →+* S is semilinear if it satisfies the two properties f (x + y) = f x + f y and f (c • x) = (σ c) • f x.

        Instances
          @[reducible, inline]
          abbrev ContinuousLinearEquivClass (F : Type u_3) (R : outParam (Type u_4)) [Semiring R] (M : outParam (Type u_5)) [TopologicalSpace M] [AddCommMonoid M] (M₂ : outParam (Type u_6)) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] [Module R M₂] [EquivLike F M M₂] :

          ContinuousLinearEquivClass F σ M M₂ asserts F is a type of bundled continuous R-linear equivs M → M₂. This is an abbreviation for ContinuousSemilinearEquivClass F (RingHom.id R) M M₂.

          Equations
          Instances For
            @[instance 100]
            instance ContinuousSemilinearEquivClass.continuousSemilinearMapClass (F : Type u_3) {R : Type u_4} {S : Type u_5} [Semiring R] [Semiring S] (σ : R →+* S) {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : Type u_6) [TopologicalSpace M] [AddCommMonoid M] (M₂ : Type u_7) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] [Module S M₂] [EquivLike F M M₂] [s : ContinuousSemilinearEquivClass F σ M M₂] :
            @[instance 100]
            instance ContinuousSemilinearEquivClass.instHomeomorphClass (F : Type u_3) {R : Type u_4} {S : Type u_5} [Semiring R] [Semiring S] (σ : R →+* S) {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : Type u_6) [TopologicalSpace M] [AddCommMonoid M] (M₂ : Type u_7) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] [Module S M₂] [EquivLike F M M₂] [s : ContinuousSemilinearEquivClass F σ M M₂] :
            def ContinuousLinearMap.iInfKerProjEquiv (R : Type u_3) [Semiring R] {ι : Type u_6} (φ : ιType u_7) [(i : ι) → TopologicalSpace (φ i)] [(i : ι) → AddCommMonoid (φ i)] [(i : ι) → Module R (φ i)] {I J : Set ι} [DecidablePred fun (i : ι) => i I] (hd : Disjoint I J) (hu : Set.univ I J) :
            (⨅ iJ, LinearMap.ker (ContinuousLinearMap.proj i)) ≃L[R] (i : I) → φ i

            If I and J are complementary index sets, the product of the kernels of the Jth projections of φ is linearly equivalent to the product over I.

            Equations
            Instances For
              def ContinuousLinearEquiv.toContinuousLinearMap {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
              M₁ →SL[σ₁₂] M₂

              A continuous linear equivalence induces a continuous linear map.

              Equations
              • e = { toLinearMap := e.toLinearEquiv, cont := }
              Instances For
                instance ContinuousLinearEquiv.ContinuousLinearMap.coe {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] :
                Coe (M₁ ≃SL[σ₁₂] M₂) (M₁ →SL[σ₁₂] M₂)

                Coerce continuous linear equivs to continuous linear maps.

                Equations
                • ContinuousLinearEquiv.ContinuousLinearMap.coe = { coe := ContinuousLinearEquiv.toContinuousLinearMap }
                instance ContinuousLinearEquiv.equivLike {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] :
                EquivLike (M₁ ≃SL[σ₁₂] M₂) M₁ M₂
                Equations
                • One or more equations did not get rendered due to their size.
                instance ContinuousLinearEquiv.continuousSemilinearEquivClass {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] :
                ContinuousSemilinearEquivClass (M₁ ≃SL[σ₁₂] M₂) σ₁₂ M₁ M₂
                theorem ContinuousLinearEquiv.coe_apply {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (b : M₁) :
                e b = e b
                @[simp]
                theorem ContinuousLinearEquiv.coe_toLinearEquiv {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (f : M₁ ≃SL[σ₁₂] M₂) :
                f.toLinearEquiv = f
                @[simp]
                theorem ContinuousLinearEquiv.coe_coe {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                e = e
                theorem ContinuousLinearEquiv.toLinearEquiv_injective {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] :
                Function.Injective ContinuousLinearEquiv.toLinearEquiv
                theorem ContinuousLinearEquiv.ext {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] {f g : M₁ ≃SL[σ₁₂] M₂} (h : f = g) :
                f = g
                theorem ContinuousLinearEquiv.coe_injective {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] :
                Function.Injective ContinuousLinearEquiv.toContinuousLinearMap
                @[simp]
                theorem ContinuousLinearEquiv.coe_inj {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] {e e' : M₁ ≃SL[σ₁₂] M₂} :
                e = e' e = e'
                def ContinuousLinearEquiv.toHomeomorph {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                M₁ ≃ₜ M₂

                A continuous linear equivalence induces a homeomorphism.

                Equations
                • e.toHomeomorph = { toEquiv := e.toEquiv, continuous_toFun := , continuous_invFun := }
                Instances For
                  @[simp]
                  theorem ContinuousLinearEquiv.coe_toHomeomorph {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                  e.toHomeomorph = e
                  theorem ContinuousLinearEquiv.isOpenMap {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                  theorem ContinuousLinearEquiv.image_closure {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) :
                  e '' closure s = closure (e '' s)
                  theorem ContinuousLinearEquiv.preimage_closure {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) :
                  e ⁻¹' closure s = closure (e ⁻¹' s)
                  @[simp]
                  theorem ContinuousLinearEquiv.isClosed_image {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) {s : Set M₁} :
                  IsClosed (e '' s) IsClosed s
                  theorem ContinuousLinearEquiv.map_nhds_eq {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) :
                  Filter.map (⇑e) (nhds x) = nhds (e x)
                  theorem ContinuousLinearEquiv.map_zero {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                  e 0 = 0
                  theorem ContinuousLinearEquiv.map_add {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (x y : M₁) :
                  e (x + y) = e x + e y
                  @[simp]
                  theorem ContinuousLinearEquiv.map_smulₛₗ {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (c : R₁) (x : M₁) :
                  e (c x) = σ₁₂ c e x
                  theorem ContinuousLinearEquiv.map_smul {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₁ M₂] (e : M₁ ≃L[R₁] M₂) (c : R₁) (x : M₁) :
                  e (c x) = c e x
                  theorem ContinuousLinearEquiv.map_eq_zero_iff {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) {x : M₁} :
                  e x = 0 x = 0
                  theorem ContinuousLinearEquiv.continuous {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                  theorem ContinuousLinearEquiv.continuousOn {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) {s : Set M₁} :
                  ContinuousOn (⇑e) s
                  theorem ContinuousLinearEquiv.continuousAt {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) {x : M₁} :
                  ContinuousAt (⇑e) x
                  theorem ContinuousLinearEquiv.continuousWithinAt {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) {s : Set M₁} {x : M₁} :
                  theorem ContinuousLinearEquiv.comp_continuousOn_iff {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] {α : Type u_10} [TopologicalSpace α] (e : M₁ ≃SL[σ₁₂] M₂) {f : αM₁} {s : Set α} :
                  theorem ContinuousLinearEquiv.comp_continuous_iff {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] {α : Type u_10} [TopologicalSpace α] (e : M₁ ≃SL[σ₁₂] M₂) {f : αM₁} :
                  theorem ContinuousLinearEquiv.ext₁ {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] [Module R₁ M₁] [TopologicalSpace R₁] {f g : R₁ ≃L[R₁] M₁} (h : f 1 = g 1) :
                  f = g

                  An extensionality lemma for R ≃L[R] M.

                  def ContinuousLinearEquiv.refl (R₁ : Type u_3) [Semiring R₁] (M₁ : Type u_6) [TopologicalSpace M₁] [AddCommMonoid M₁] [Module R₁ M₁] :
                  M₁ ≃L[R₁] M₁

                  The identity map as a continuous linear equivalence.

                  Equations
                  Instances For
                    @[simp]
                    theorem ContinuousLinearEquiv.refl_apply (R₁ : Type u_3) [Semiring R₁] (M₁ : Type u_6) [TopologicalSpace M₁] [AddCommMonoid M₁] [Module R₁ M₁] (x : M₁) :
                    @[simp]
                    theorem ContinuousLinearEquiv.coe_refl {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] [Module R₁ M₁] :
                    @[simp]
                    theorem ContinuousLinearEquiv.coe_refl' {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] [Module R₁ M₁] :
                    (ContinuousLinearEquiv.refl R₁ M₁) = id
                    def ContinuousLinearEquiv.symm {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                    M₂ ≃SL[σ₂₁] M₁

                    The inverse of a continuous linear equivalence as a continuous linear equivalence

                    Equations
                    • e.symm = { toLinearEquiv := e.symm, continuous_toFun := , continuous_invFun := }
                    Instances For
                      @[simp]
                      theorem ContinuousLinearEquiv.symm_toLinearEquiv {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                      e.symm.toLinearEquiv = e.symm
                      @[simp]
                      theorem ContinuousLinearEquiv.symm_toHomeomorph {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                      e.toHomeomorph.symm = e.symm.toHomeomorph
                      def ContinuousLinearEquiv.Simps.apply {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (h : M₁ ≃SL[σ₁₂] M₂) :
                      M₁M₂

                      See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections.

                      Equations
                      Instances For
                        def ContinuousLinearEquiv.Simps.symm_apply {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (h : M₁ ≃SL[σ₁₂] M₂) :
                        M₂M₁

                        See Note [custom simps projection]

                        Equations
                        Instances For
                          theorem ContinuousLinearEquiv.symm_map_nhds_eq {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) :
                          Filter.map (⇑e.symm) (nhds (e x)) = nhds x
                          def ContinuousLinearEquiv.trans {R₁ : Type u_3} {R₂ : Type u_4} {R₃ : Type u_5} [Semiring R₁] [Semiring R₂] [Semiring R₃] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] [Module R₁ M₁] [Module R₂ M₂] [Module R₃ M₃] (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃SL[σ₂₃] M₃) :
                          M₁ ≃SL[σ₁₃] M₃

                          The composition of two continuous linear equivalences as a continuous linear equivalence.

                          Equations
                          • e₁.trans e₂ = { toLinearEquiv := e₁.trans e₂.toLinearEquiv, continuous_toFun := , continuous_invFun := }
                          Instances For
                            @[simp]
                            theorem ContinuousLinearEquiv.trans_toLinearEquiv {R₁ : Type u_3} {R₂ : Type u_4} {R₃ : Type u_5} [Semiring R₁] [Semiring R₂] [Semiring R₃] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] [Module R₁ M₁] [Module R₂ M₂] [Module R₃ M₃] (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃SL[σ₂₃] M₃) :
                            (e₁.trans e₂).toLinearEquiv = e₁.trans e₂.toLinearEquiv
                            def ContinuousLinearEquiv.prod {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] {M₄ : Type u_9} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) (e' : M₃ ≃L[R₁] M₄) :
                            (M₁ × M₃) ≃L[R₁] M₂ × M₄

                            Product of two continuous linear equivalences. The map comes from Equiv.prodCongr.

                            Equations
                            • e.prod e' = { toLinearEquiv := e.prod e'.toLinearEquiv, continuous_toFun := , continuous_invFun := }
                            Instances For
                              @[simp]
                              theorem ContinuousLinearEquiv.prod_apply {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] {M₄ : Type u_9} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) (e' : M₃ ≃L[R₁] M₄) (x : M₁ × M₃) :
                              (e.prod e') x = (e x.1, e' x.2)
                              @[simp]
                              theorem ContinuousLinearEquiv.coe_prod {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] {M₄ : Type u_9} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) (e' : M₃ ≃L[R₁] M₄) :
                              (e.prod e') = (↑e).prodMap e'
                              theorem ContinuousLinearEquiv.prod_symm {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] {M₄ : Type u_9} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) (e' : M₃ ≃L[R₁] M₄) :
                              (e.prod e').symm = e.symm.prod e'.symm
                              def ContinuousLinearEquiv.prodComm (R₁ : Type u_3) [Semiring R₁] (M₁ : Type u_6) [TopologicalSpace M₁] [AddCommMonoid M₁] (M₂ : Type u_7) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₁ M₂] :
                              (M₁ × M₂) ≃L[R₁] M₂ × M₁

                              Product of modules is commutative up to continuous linear isomorphism.

                              Equations
                              Instances For
                                @[simp]
                                theorem ContinuousLinearEquiv.prodComm_toLinearEquiv (R₁ : Type u_3) [Semiring R₁] (M₁ : Type u_6) [TopologicalSpace M₁] [AddCommMonoid M₁] (M₂ : Type u_7) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₁ M₂] :
                                (ContinuousLinearEquiv.prodComm R₁ M₁ M₂).toLinearEquiv = LinearEquiv.prodComm R₁ M₁ M₂
                                @[simp]
                                theorem ContinuousLinearEquiv.prodComm_apply (R₁ : Type u_3) [Semiring R₁] (M₁ : Type u_6) [TopologicalSpace M₁] [AddCommMonoid M₁] (M₂ : Type u_7) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₁ M₂] (a✝ : M₁ × M₂) :
                                (ContinuousLinearEquiv.prodComm R₁ M₁ M₂) a✝ = a✝.swap
                                @[simp]
                                theorem ContinuousLinearEquiv.prodComm_symm (R₁ : Type u_3) [Semiring R₁] (M₁ : Type u_6) [TopologicalSpace M₁] [AddCommMonoid M₁] (M₂ : Type u_7) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₁ M₂] :
                                theorem ContinuousLinearEquiv.bijective {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                                theorem ContinuousLinearEquiv.injective {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                                theorem ContinuousLinearEquiv.surjective {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                                @[simp]
                                theorem ContinuousLinearEquiv.trans_apply {R₁ : Type u_3} {R₂ : Type u_4} {R₃ : Type u_5} [Semiring R₁] [Semiring R₂] [Semiring R₃] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] [Module R₁ M₁] [Module R₂ M₂] [Module R₃ M₃] (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃SL[σ₂₃] M₃) (c : M₁) :
                                (e₁.trans e₂) c = e₂ (e₁ c)
                                @[simp]
                                theorem ContinuousLinearEquiv.apply_symm_apply {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (c : M₂) :
                                e (e.symm c) = c
                                @[simp]
                                theorem ContinuousLinearEquiv.symm_apply_apply {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (b : M₁) :
                                e.symm (e b) = b
                                @[simp]
                                theorem ContinuousLinearEquiv.symm_trans_apply {R₁ : Type u_3} {R₂ : Type u_4} {R₃ : Type u_5} [Semiring R₁] [Semiring R₂] [Semiring R₃] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] [Module R₁ M₁] [Module R₂ M₂] [Module R₃ M₃] (e₁ : M₂ ≃SL[σ₂₁] M₁) (e₂ : M₃ ≃SL[σ₃₂] M₂) (c : M₁) :
                                (e₂.trans e₁).symm c = e₂.symm (e₁.symm c)
                                @[simp]
                                theorem ContinuousLinearEquiv.symm_image_image {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) :
                                e.symm '' (e '' s) = s
                                @[simp]
                                theorem ContinuousLinearEquiv.image_symm_image {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) :
                                e '' (e.symm '' s) = s
                                @[simp]
                                theorem ContinuousLinearEquiv.comp_coe {R₁ : Type u_3} {R₂ : Type u_4} {R₃ : Type u_5} [Semiring R₁] [Semiring R₂] [Semiring R₃] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] [Module R₁ M₁] [Module R₂ M₂] [Module R₃ M₃] (f : M₁ ≃SL[σ₁₂] M₂) (f' : M₂ ≃SL[σ₂₃] M₃) :
                                (↑f').comp f = (f.trans f')
                                @[simp]
                                theorem ContinuousLinearEquiv.coe_comp_coe_symm {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                                (↑e).comp e.symm = ContinuousLinearMap.id R₂ M₂
                                @[simp]
                                theorem ContinuousLinearEquiv.coe_symm_comp_coe {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                                (↑e.symm).comp e = ContinuousLinearMap.id R₁ M₁
                                @[simp]
                                theorem ContinuousLinearEquiv.symm_comp_self {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                                e.symm e = id
                                @[simp]
                                theorem ContinuousLinearEquiv.self_comp_symm {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                                e e.symm = id
                                @[simp]
                                theorem ContinuousLinearEquiv.symm_symm {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) :
                                e.symm.symm = e
                                @[simp]
                                theorem ContinuousLinearEquiv.refl_symm {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] [Module R₁ M₁] :
                                theorem ContinuousLinearEquiv.symm_symm_apply {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) :
                                e.symm.symm x = e x
                                theorem ContinuousLinearEquiv.symm_apply_eq {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) {x : M₂} {y : M₁} :
                                e.symm x = y x = e y
                                theorem ContinuousLinearEquiv.eq_symm_apply {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) {x : M₂} {y : M₁} :
                                y = e.symm x e y = x
                                theorem ContinuousLinearEquiv.image_eq_preimage {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) :
                                e '' s = e.symm ⁻¹' s
                                theorem ContinuousLinearEquiv.image_symm_eq_preimage {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) :
                                e.symm '' s = e ⁻¹' s
                                @[simp]
                                theorem ContinuousLinearEquiv.symm_preimage_preimage {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) :
                                e.symm ⁻¹' (e ⁻¹' s) = s
                                @[simp]
                                theorem ContinuousLinearEquiv.preimage_symm_preimage {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) :
                                e ⁻¹' (e.symm ⁻¹' s) = s
                                theorem ContinuousLinearEquiv.isUniformEmbedding {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {E₁ : Type u_10} {E₂ : Type u_11} [UniformSpace E₁] [UniformSpace E₂] [AddCommGroup E₁] [AddCommGroup E₂] [Module R₁ E₁] [Module R₂ E₂] [UniformAddGroup E₁] [UniformAddGroup E₂] (e : E₁ ≃SL[σ₁₂] E₂) :
                                @[deprecated ContinuousLinearEquiv.isUniformEmbedding]
                                theorem ContinuousLinearEquiv.uniformEmbedding {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {E₁ : Type u_10} {E₂ : Type u_11} [UniformSpace E₁] [UniformSpace E₂] [AddCommGroup E₁] [AddCommGroup E₂] [Module R₁ E₁] [Module R₂ E₂] [UniformAddGroup E₁] [UniformAddGroup E₂] (e : E₁ ≃SL[σ₁₂] E₂) :

                                Alias of ContinuousLinearEquiv.isUniformEmbedding.

                                theorem LinearEquiv.isUniformEmbedding {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {E₁ : Type u_10} {E₂ : Type u_11} [UniformSpace E₁] [UniformSpace E₂] [AddCommGroup E₁] [AddCommGroup E₂] [Module R₁ E₁] [Module R₂ E₂] [UniformAddGroup E₁] [UniformAddGroup E₂] (e : E₁ ≃ₛₗ[σ₁₂] E₂) (h₁ : Continuous e) (h₂ : Continuous e.symm) :
                                @[deprecated LinearEquiv.isUniformEmbedding]
                                theorem LinearEquiv.uniformEmbedding {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {E₁ : Type u_10} {E₂ : Type u_11} [UniformSpace E₁] [UniformSpace E₂] [AddCommGroup E₁] [AddCommGroup E₂] [Module R₁ E₁] [Module R₂ E₂] [UniformAddGroup E₁] [UniformAddGroup E₂] (e : E₁ ≃ₛₗ[σ₁₂] E₂) (h₁ : Continuous e) (h₂ : Continuous e.symm) :

                                Alias of LinearEquiv.isUniformEmbedding.

                                def ContinuousLinearEquiv.equivOfInverse {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) (h₁ : Function.LeftInverse f₂ f₁) (h₂ : Function.RightInverse f₂ f₁) :
                                M₁ ≃SL[σ₁₂] M₂

                                Create a ContinuousLinearEquiv from two ContinuousLinearMaps that are inverse of each other. See also equivOfInverse'.

                                Equations
                                • ContinuousLinearEquiv.equivOfInverse f₁ f₂ h₁ h₂ = { toLinearMap := f₁, invFun := f₂, left_inv := h₁, right_inv := h₂, continuous_toFun := , continuous_invFun := }
                                Instances For
                                  @[simp]
                                  theorem ContinuousLinearEquiv.equivOfInverse_apply {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) (h₁ : Function.LeftInverse f₂ f₁) (h₂ : Function.RightInverse f₂ f₁) (x : M₁) :
                                  (ContinuousLinearEquiv.equivOfInverse f₁ f₂ h₁ h₂) x = f₁ x
                                  @[simp]
                                  theorem ContinuousLinearEquiv.symm_equivOfInverse {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) (h₁ : Function.LeftInverse f₂ f₁) (h₂ : Function.RightInverse f₂ f₁) :
                                  (ContinuousLinearEquiv.equivOfInverse f₁ f₂ h₁ h₂).symm = ContinuousLinearEquiv.equivOfInverse f₂ f₁ h₂ h₁
                                  def ContinuousLinearEquiv.equivOfInverse' {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) (h₁ : f₁.comp f₂ = ContinuousLinearMap.id R₂ M₂) (h₂ : f₂.comp f₁ = ContinuousLinearMap.id R₁ M₁) :
                                  M₁ ≃SL[σ₁₂] M₂

                                  Create a ContinuousLinearEquiv from two ContinuousLinearMaps that are inverse of each other, in the ContinuousLinearMap.comp sense. See also equivOfInverse.

                                  Equations
                                  Instances For
                                    @[simp]
                                    theorem ContinuousLinearEquiv.equivOfInverse'_apply {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) (h₁ : f₁.comp f₂ = ContinuousLinearMap.id R₂ M₂) (h₂ : f₂.comp f₁ = ContinuousLinearMap.id R₁ M₁) (x : M₁) :
                                    (ContinuousLinearEquiv.equivOfInverse' f₁ f₂ h₁ h₂) x = f₁ x
                                    @[simp]
                                    theorem ContinuousLinearEquiv.symm_equivOfInverse' {R₁ : Type u_3} {R₂ : Type u_4} [Semiring R₁] [Semiring R₂] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₁ M₁] [Module R₂ M₂] (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) (h₁ : f₁.comp f₂ = ContinuousLinearMap.id R₂ M₂) (h₂ : f₂.comp f₁ = ContinuousLinearMap.id R₁ M₁) :

                                    The inverse of equivOfInverse' is obtained by swapping the order of its parameters.

                                    instance ContinuousLinearEquiv.automorphismGroup {R₁ : Type u_3} [Semiring R₁] (M₁ : Type u_6) [TopologicalSpace M₁] [AddCommMonoid M₁] [Module R₁ M₁] :
                                    Group (M₁ ≃L[R₁] M₁)

                                    The continuous linear equivalences from M to itself form a group under composition.

                                    Equations
                                    def ContinuousLinearEquiv.ulift {R₁ : Type u_3} [Semiring R₁] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] [Module R₁ M₁] :

                                    The continuous linear equivalence between ULift M₁ and M₁.

                                    This is a continuous version of ULift.moduleEquiv.

                                    Equations
                                    • ContinuousLinearEquiv.ulift = { toLinearEquiv := ULift.moduleEquiv, continuous_toFun := , continuous_invFun := }
                                    Instances For
                                      def ContinuousLinearEquiv.arrowCongrEquiv {R₁ : Type u_3} {R₂ : Type u_4} {R₃ : Type u_5} [Semiring R₁] [Semiring R₂] [Semiring R₃] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R₁ →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] {M₄ : Type u_9} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] [Module R₂ M₂] [Module R₃ M₃] {R₄ : Type u_10} [Semiring R₄] [Module R₄ M₄] {σ₃₄ : R₃ →+* R₄} {σ₄₃ : R₄ →+* R₃} [RingHomInvPair σ₃₄ σ₄₃] [RingHomInvPair σ₄₃ σ₃₄] {σ₂₄ : R₂ →+* R₄} {σ₁₄ : R₁ →+* R₄} [RingHomCompTriple σ₂₁ σ₁₄ σ₂₄] [RingHomCompTriple σ₂₄ σ₄₃ σ₂₃] [RingHomCompTriple σ₁₃ σ₃₄ σ₁₄] (e₁₂ : M₁ ≃SL[σ₁₂] M₂) (e₄₃ : M₄ ≃SL[σ₄₃] M₃) :
                                      (M₁ →SL[σ₁₄] M₄) (M₂ →SL[σ₂₃] M₃)

                                      A pair of continuous (semi)linear equivalences generates an equivalence between the spaces of continuous linear maps. See also ContinuousLinearEquiv.arrowCongr.

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        @[simp]
                                        theorem ContinuousLinearEquiv.arrowCongrEquiv_symm_apply {R₁ : Type u_3} {R₂ : Type u_4} {R₃ : Type u_5} [Semiring R₁] [Semiring R₂] [Semiring R₃] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R₁ →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] {M₄ : Type u_9} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] [Module R₂ M₂] [Module R₃ M₃] {R₄ : Type u_10} [Semiring R₄] [Module R₄ M₄] {σ₃₄ : R₃ →+* R₄} {σ₄₃ : R₄ →+* R₃} [RingHomInvPair σ₃₄ σ₄₃] [RingHomInvPair σ₄₃ σ₃₄] {σ₂₄ : R₂ →+* R₄} {σ₁₄ : R₁ →+* R₄} [RingHomCompTriple σ₂₁ σ₁₄ σ₂₄] [RingHomCompTriple σ₂₄ σ₄₃ σ₂₃] [RingHomCompTriple σ₁₃ σ₃₄ σ₁₄] (e₁₂ : M₁ ≃SL[σ₁₂] M₂) (e₄₃ : M₄ ≃SL[σ₄₃] M₃) (f : M₂ →SL[σ₂₃] M₃) :
                                        (e₁₂.arrowCongrEquiv e₄₃).symm f = (↑e₄₃.symm).comp (f.comp e₁₂)
                                        @[simp]
                                        theorem ContinuousLinearEquiv.arrowCongrEquiv_apply {R₁ : Type u_3} {R₂ : Type u_4} {R₃ : Type u_5} [Semiring R₁] [Semiring R₂] [Semiring R₃] {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R₁ →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] {M₁ : Type u_6} [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type u_7} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type u_8} [TopologicalSpace M₃] [AddCommMonoid M₃] {M₄ : Type u_9} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] [Module R₂ M₂] [Module R₃ M₃] {R₄ : Type u_10} [Semiring R₄] [Module R₄ M₄] {σ₃₄ : R₃ →+* R₄} {σ₄₃ : R₄ →+* R₃} [RingHomInvPair σ₃₄ σ₄₃] [RingHomInvPair σ₄₃ σ₃₄] {σ₂₄ : R₂ →+* R₄} {σ₁₄ : R₁ →+* R₄} [RingHomCompTriple σ₂₁ σ₁₄ σ₂₄] [RingHomCompTriple σ₂₄ σ₄₃ σ₂₃] [RingHomCompTriple σ₁₃ σ₃₄ σ₁₄] (e₁₂ : M₁ ≃SL[σ₁₂] M₂) (e₄₃ : M₄ ≃SL[σ₄₃] M₃) (f : M₁ →SL[σ₁₄] M₄) :
                                        (e₁₂.arrowCongrEquiv e₄₃) f = (↑e₄₃).comp (f.comp e₁₂.symm)
                                        def ContinuousLinearEquiv.piCongrRight {R₁ : Type u_3} [Semiring R₁] {ι : Type u_11} {M : ιType u_12} [(i : ι) → TopologicalSpace (M i)] [(i : ι) → AddCommMonoid (M i)] [(i : ι) → Module R₁ (M i)] {N : ιType u_13} [(i : ι) → TopologicalSpace (N i)] [(i : ι) → AddCommMonoid (N i)] [(i : ι) → Module R₁ (N i)] (f : (i : ι) → M i ≃L[R₁] N i) :
                                        ((i : ι) → M i) ≃L[R₁] (i : ι) → N i

                                        Combine a family of continuous linear equivalences into a continuous linear equivalence of pi-types.

                                        Equations
                                        Instances For
                                          @[simp]
                                          theorem ContinuousLinearEquiv.piCongrRight_apply {R₁ : Type u_3} [Semiring R₁] {ι : Type u_11} {M : ιType u_12} [(i : ι) → TopologicalSpace (M i)] [(i : ι) → AddCommMonoid (M i)] [(i : ι) → Module R₁ (M i)] {N : ιType u_13} [(i : ι) → TopologicalSpace (N i)] [(i : ι) → AddCommMonoid (N i)] [(i : ι) → Module R₁ (N i)] (f : (i : ι) → M i ≃L[R₁] N i) (m : (i : ι) → M i) (i : ι) :
                                          @[simp]
                                          theorem ContinuousLinearEquiv.piCongrRight_symm_apply {R₁ : Type u_3} [Semiring R₁] {ι : Type u_11} {M : ιType u_12} [(i : ι) → TopologicalSpace (M i)] [(i : ι) → AddCommMonoid (M i)] [(i : ι) → Module R₁ (M i)] {N : ιType u_13} [(i : ι) → TopologicalSpace (N i)] [(i : ι) → AddCommMonoid (N i)] [(i : ι) → Module R₁ (N i)] (f : (i : ι) → M i ≃L[R₁] N i) (n : (i : ι) → N i) (i : ι) :
                                          (ContinuousLinearEquiv.piCongrRight f).symm n i = (f i).symm (n i)
                                          def ContinuousLinearEquiv.skewProd {R : Type u_3} [Semiring R] {M : Type u_4} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type u_5} [TopologicalSpace M₂] [AddCommGroup M₂] {M₃ : Type u_6} [TopologicalSpace M₃] [AddCommGroup M₃] {M₄ : Type u_7} [TopologicalSpace M₄] [AddCommGroup M₄] [Module R M] [Module R M₂] [Module R M₃] [Module R M₄] [TopologicalAddGroup M₄] (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M →L[R] M₄) :
                                          (M × M₃) ≃L[R] M₂ × M₄

                                          Equivalence given by a block lower diagonal matrix. e and e' are diagonal square blocks, and f is a rectangular block below the diagonal.

                                          Equations
                                          • e.skewProd e' f = { toLinearEquiv := e.skewProd e'.toLinearEquiv f, continuous_toFun := , continuous_invFun := }
                                          Instances For
                                            @[simp]
                                            theorem ContinuousLinearEquiv.skewProd_apply {R : Type u_3} [Semiring R] {M : Type u_4} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type u_5} [TopologicalSpace M₂] [AddCommGroup M₂] {M₃ : Type u_6} [TopologicalSpace M₃] [AddCommGroup M₃] {M₄ : Type u_7} [TopologicalSpace M₄] [AddCommGroup M₄] [Module R M] [Module R M₂] [Module R M₃] [Module R M₄] [TopologicalAddGroup M₄] (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M →L[R] M₄) (x : M × M₃) :
                                            (e.skewProd e' f) x = (e x.1, e' x.2 + f x.1)
                                            @[simp]
                                            theorem ContinuousLinearEquiv.skewProd_symm_apply {R : Type u_3} [Semiring R] {M : Type u_4} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type u_5} [TopologicalSpace M₂] [AddCommGroup M₂] {M₃ : Type u_6} [TopologicalSpace M₃] [AddCommGroup M₃] {M₄ : Type u_7} [TopologicalSpace M₄] [AddCommGroup M₄] [Module R M] [Module R M₂] [Module R M₃] [Module R M₄] [TopologicalAddGroup M₄] (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M →L[R] M₄) (x : M₂ × M₄) :
                                            (e.skewProd e' f).symm x = (e.symm x.1, e'.symm (x.2 - f (e.symm x.1)))

                                            The negation map as a continuous linear equivalence.

                                            Equations
                                            Instances For
                                              theorem ContinuousLinearEquiv.map_sub {R : Type u_3} [Ring R] {R₂ : Type u_4} [Ring R₂] {M : Type u_5} [TopologicalSpace M] [AddCommGroup M] [Module R M] {M₂ : Type u_6} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R₂ M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] (e : M ≃SL[σ₁₂] M₂) (x y : M) :
                                              e (x - y) = e x - e y
                                              theorem ContinuousLinearEquiv.map_neg {R : Type u_3} [Ring R] {R₂ : Type u_4} [Ring R₂] {M : Type u_5} [TopologicalSpace M] [AddCommGroup M] [Module R M] {M₂ : Type u_6} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R₂ M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] (e : M ≃SL[σ₁₂] M₂) (x : M) :
                                              e (-x) = -e x

                                              The next theorems cover the identification between M ≃L[R] Mand the group of units of the ring M →L[R] M.

                                              def ContinuousLinearEquiv.ofUnit {R : Type u_3} [Ring R] {M : Type u_5} [TopologicalSpace M] [AddCommGroup M] [Module R M] (f : (M →L[R] M)ˣ) :
                                              M ≃L[R] M

                                              An invertible continuous linear map f determines a continuous equivalence from M to itself.

                                              Equations
                                              • ContinuousLinearEquiv.ofUnit f = { toFun := f, map_add' := , map_smul' := , invFun := f.inv, left_inv := , right_inv := , continuous_toFun := , continuous_invFun := }
                                              Instances For
                                                def ContinuousLinearEquiv.toUnit {R : Type u_3} [Ring R] {M : Type u_5} [TopologicalSpace M] [AddCommGroup M] [Module R M] (f : M ≃L[R] M) :
                                                (M →L[R] M)ˣ

                                                A continuous equivalence from M to itself determines an invertible continuous linear map.

                                                Equations
                                                • f.toUnit = { val := f, inv := f.symm, val_inv := , inv_val := }
                                                Instances For
                                                  def ContinuousLinearEquiv.unitsEquiv (R : Type u_3) [Ring R] (M : Type u_5) [TopologicalSpace M] [AddCommGroup M] [Module R M] :
                                                  (M →L[R] M)ˣ ≃* M ≃L[R] M

                                                  The units of the algebra of continuous R-linear endomorphisms of M is multiplicatively equivalent to the type of continuous linear equivalences between M and itself.

                                                  Equations
                                                  • ContinuousLinearEquiv.unitsEquiv R M = { toFun := ContinuousLinearEquiv.ofUnit, invFun := ContinuousLinearEquiv.toUnit, left_inv := , right_inv := , map_mul' := }
                                                  Instances For
                                                    @[simp]
                                                    theorem ContinuousLinearEquiv.unitsEquiv_apply (R : Type u_3) [Ring R] (M : Type u_5) [TopologicalSpace M] [AddCommGroup M] [Module R M] (f : (M →L[R] M)ˣ) (x : M) :

                                                    Continuous linear equivalences R ≃L[R] R are enumerated by .

                                                    Equations
                                                    • One or more equations did not get rendered due to their size.
                                                    Instances For
                                                      def ContinuousLinearEquiv.equivOfRightInverse {R : Type u_3} [Ring R] {M : Type u_5} [TopologicalSpace M] [AddCommGroup M] [Module R M] {M₂ : Type u_6} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R M₂] [TopologicalAddGroup M] (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) (h : Function.RightInverse f₂ f₁) :
                                                      M ≃L[R] M₂ × (LinearMap.ker f₁)

                                                      A pair of continuous linear maps such that f₁ ∘ f₂ = id generates a continuous linear equivalence e between M and M₂ × f₁.ker such that (e x).2 = x for x ∈ f₁.ker, (e x).1 = f₁ x, and (e (f₂ y)).2 = 0. The map is given by e x = (f₁ x, x - f₂ (f₁ x)).

                                                      Equations
                                                      Instances For
                                                        @[simp]
                                                        theorem ContinuousLinearEquiv.fst_equivOfRightInverse {R : Type u_3} [Ring R] {M : Type u_5} [TopologicalSpace M] [AddCommGroup M] [Module R M] {M₂ : Type u_6} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R M₂] [TopologicalAddGroup M] (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) (h : Function.RightInverse f₂ f₁) (x : M) :
                                                        ((ContinuousLinearEquiv.equivOfRightInverse f₁ f₂ h) x).1 = f₁ x
                                                        @[simp]
                                                        theorem ContinuousLinearEquiv.snd_equivOfRightInverse {R : Type u_3} [Ring R] {M : Type u_5} [TopologicalSpace M] [AddCommGroup M] [Module R M] {M₂ : Type u_6} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R M₂] [TopologicalAddGroup M] (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) (h : Function.RightInverse f₂ f₁) (x : M) :
                                                        ((ContinuousLinearEquiv.equivOfRightInverse f₁ f₂ h) x).2 = x - f₂ (f₁ x)
                                                        @[simp]
                                                        theorem ContinuousLinearEquiv.equivOfRightInverse_symm_apply {R : Type u_3} [Ring R] {M : Type u_5} [TopologicalSpace M] [AddCommGroup M] [Module R M] {M₂ : Type u_6} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R M₂] [TopologicalAddGroup M] (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) (h : Function.RightInverse f₂ f₁) (y : M₂ × (LinearMap.ker f₁)) :
                                                        (ContinuousLinearEquiv.equivOfRightInverse f₁ f₂ h).symm y = f₂ y.1 + y.2
                                                        def ContinuousLinearEquiv.funUnique (ι : Type u_3) (R : Type u_4) (M : Type u_5) [Unique ι] [Semiring R] [AddCommMonoid M] [Module R M] [TopologicalSpace M] :
                                                        (ιM) ≃L[R] M

                                                        If ι has a unique element, then ι → M is continuously linear equivalent to M.

                                                        Equations
                                                        Instances For
                                                          @[simp]
                                                          theorem ContinuousLinearEquiv.coe_funUnique {ι : Type u_3} {R : Type u_4} {M : Type u_5} [Unique ι] [Semiring R] [AddCommMonoid M] [Module R M] [TopologicalSpace M] :
                                                          @[simp]
                                                          def ContinuousLinearEquiv.piFinTwo (R : Type u_4) [Semiring R] (M : Fin 2Type u_6) [(i : Fin 2) → AddCommMonoid (M i)] [(i : Fin 2) → Module R (M i)] [(i : Fin 2) → TopologicalSpace (M i)] :
                                                          ((i : Fin 2) → M i) ≃L[R] M 0 × M 1

                                                          Continuous linear equivalence between dependent functions (i : Fin 2) → M i and M 0 × M 1.

                                                          Equations
                                                          Instances For
                                                            @[simp]
                                                            theorem ContinuousLinearEquiv.piFinTwo_symm_apply (R : Type u_4) [Semiring R] (M : Fin 2Type u_6) [(i : Fin 2) → AddCommMonoid (M i)] [(i : Fin 2) → Module R (M i)] [(i : Fin 2) → TopologicalSpace (M i)] :
                                                            (ContinuousLinearEquiv.piFinTwo R M).symm = fun (p : M 0 × M 1) => Fin.cons p.1 (Fin.cons p.2 finZeroElim)
                                                            @[simp]
                                                            theorem ContinuousLinearEquiv.piFinTwo_apply (R : Type u_4) [Semiring R] (M : Fin 2Type u_6) [(i : Fin 2) → AddCommMonoid (M i)] [(i : Fin 2) → Module R (M i)] [(i : Fin 2) → TopologicalSpace (M i)] :
                                                            (ContinuousLinearEquiv.piFinTwo R M) = fun (f : (i : Fin 2) → M i) => (f 0, f 1)
                                                            def ContinuousLinearEquiv.finTwoArrow (R : Type u_4) (M : Type u_5) [Semiring R] [AddCommMonoid M] [Module R M] [TopologicalSpace M] :
                                                            (Fin 2M) ≃L[R] M × M

                                                            Continuous linear equivalence between vectors in M² = Fin 2 → M and M × M.

                                                            Equations
                                                            Instances For
                                                              @[simp]
                                                              theorem ContinuousLinearEquiv.finTwoArrow_apply (R : Type u_4) (M : Type u_5) [Semiring R] [AddCommMonoid M] [Module R M] [TopologicalSpace M] :
                                                              (ContinuousLinearEquiv.finTwoArrow R M) = fun (f : Fin 2M) => (f 0, f 1)
                                                              @[simp]
                                                              theorem ContinuousLinearEquiv.finTwoArrow_symm_apply (R : Type u_4) (M : Type u_5) [Semiring R] [AddCommMonoid M] [Module R M] [TopologicalSpace M] :
                                                              (ContinuousLinearEquiv.finTwoArrow R M).symm = fun (x : M × M) => ![x.1, x.2]
                                                              def ContinuousLinearMap.IsInvertible {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] (f : M →L[R] M₂) :

                                                              A continuous linear map is invertible if it is the forward direction of a continuous linear equivalence.

                                                              Equations
                                                              • f.IsInvertible = ∃ (A : M ≃L[R] M₂), A = f
                                                              Instances For
                                                                noncomputable def ContinuousLinearMap.inverse {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] :
                                                                (M →L[R] M₂)M₂ →L[R] M

                                                                Introduce a function inverse from M →L[R] M₂ to M₂ →L[R] M, which sends f to f.symm if f is a continuous linear equivalence and to 0 otherwise. This definition is somewhat ad hoc, but one needs a fully (rather than partially) defined inverse function for some purposes, including for calculus.

                                                                Equations
                                                                Instances For
                                                                  @[simp]
                                                                  theorem ContinuousLinearMap.isInvertible_equiv {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] {f : M ≃L[R] M₂} :
                                                                  (↑f).IsInvertible
                                                                  @[simp]
                                                                  theorem ContinuousLinearMap.inverse_equiv {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] (e : M ≃L[R] M₂) :
                                                                  (↑e).inverse = e.symm

                                                                  By definition, if f is invertible then inverse f = f.symm.

                                                                  @[simp]
                                                                  theorem ContinuousLinearMap.inverse_of_not_isInvertible {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] {f : M →L[R] M₂} (hf : ¬f.IsInvertible) :
                                                                  f.inverse = 0

                                                                  By definition, if f is not invertible then inverse f = 0.

                                                                  @[deprecated ContinuousLinearMap.inverse_of_not_isInvertible]
                                                                  theorem ContinuousLinearMap.inverse_non_equiv {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] {f : M →L[R] M₂} (hf : ¬f.IsInvertible) :
                                                                  f.inverse = 0

                                                                  Alias of ContinuousLinearMap.inverse_of_not_isInvertible.


                                                                  By definition, if f is not invertible then inverse f = 0.

                                                                  @[simp]
                                                                  theorem ContinuousLinearMap.IsInvertible.comp {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} {M₃ : Type u_6} [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] {g : M₂ →L[R] M₃} {f : M →L[R] M₂} (hg : g.IsInvertible) (hf : f.IsInvertible) :
                                                                  (g.comp f).IsInvertible
                                                                  theorem ContinuousLinearMap.IsInvertible.of_inverse {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] {f : M →L[R] M₂} {g : M₂ →L[R] M} (hf : f.comp g = ContinuousLinearMap.id R M₂) (hg : g.comp f = ContinuousLinearMap.id R M) :
                                                                  f.IsInvertible
                                                                  theorem ContinuousLinearMap.inverse_eq {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] {f : M →L[R] M₂} {g : M₂ →L[R] M} (hf : f.comp g = ContinuousLinearMap.id R M₂) (hg : g.comp f = ContinuousLinearMap.id R M) :
                                                                  f.inverse = g
                                                                  theorem ContinuousLinearMap.IsInvertible.inverse_apply_eq {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] {f : M →L[R] M₂} {x : M} {y : M₂} (hf : f.IsInvertible) :
                                                                  f.inverse y = x y = f x
                                                                  @[simp]
                                                                  theorem ContinuousLinearMap.isInvertible_equiv_comp {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} {M₃ : Type u_6} [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] {e : M₂ ≃L[R] M₃} {f : M →L[R] M₂} :
                                                                  ((↑e).comp f).IsInvertible f.IsInvertible
                                                                  @[simp]
                                                                  theorem ContinuousLinearMap.isInvertible_comp_equiv {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} {M₃ : Type u_6} [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] {e : M₃ ≃L[R] M} {f : M →L[R] M₂} :
                                                                  (f.comp e).IsInvertible f.IsInvertible
                                                                  @[simp]
                                                                  theorem ContinuousLinearMap.inverse_equiv_comp {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} {M₃ : Type u_6} [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] {e : M₂ ≃L[R] M₃} {f : M →L[R] M₂} :
                                                                  ((↑e).comp f).inverse = f.inverse.comp e.symm
                                                                  @[simp]
                                                                  theorem ContinuousLinearMap.inverse_comp_equiv {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} {M₃ : Type u_6} [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] {e : M₃ ≃L[R] M} {f : M →L[R] M₂} :
                                                                  (f.comp e).inverse = (↑e.symm).comp f.inverse
                                                                  theorem ContinuousLinearMap.IsInvertible.inverse_comp_of_left {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} {M₃ : Type u_6} [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] {g : M₂ →L[R] M₃} {f : M →L[R] M₂} (hg : g.IsInvertible) :
                                                                  (g.comp f).inverse = f.inverse.comp g.inverse
                                                                  theorem ContinuousLinearMap.IsInvertible.inverse_comp_apply_of_left {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} {M₃ : Type u_6} [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] {g : M₂ →L[R] M₃} {f : M →L[R] M₂} {v : M₃} (hg : g.IsInvertible) :
                                                                  (g.comp f).inverse v = f.inverse (g.inverse v)
                                                                  theorem ContinuousLinearMap.IsInvertible.inverse_comp_of_right {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} {M₃ : Type u_6} [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] {g : M₂ →L[R] M₃} {f : M →L[R] M₂} (hf : f.IsInvertible) :
                                                                  (g.comp f).inverse = f.inverse.comp g.inverse
                                                                  theorem ContinuousLinearMap.IsInvertible.inverse_comp_apply_of_right {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} {M₃ : Type u_6} [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] {g : M₂ →L[R] M₃} {f : M →L[R] M₂} {v : M₃} (hf : f.IsInvertible) :
                                                                  (g.comp f).inverse v = f.inverse (g.inverse v)
                                                                  @[simp]
                                                                  theorem ContinuousLinearMap.ring_inverse_equiv {R : Type u_3} {M : Type u_4} [TopologicalSpace M] [Ring R] [AddCommGroup M] [Module R M] (e : M ≃L[R] M) :
                                                                  Ring.inverse e = (↑e).inverse
                                                                  theorem ContinuousLinearMap.to_ring_inverse {R : Type u_3} {M : Type u_4} {M₂ : Type u_5} [TopologicalSpace M] [TopologicalSpace M₂] [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup M₂] [Module R M₂] (e : M ≃L[R] M₂) (f : M →L[R] M₂) :
                                                                  f.inverse = (Ring.inverse ((↑e.symm).comp f)).comp e.symm

                                                                  The function ContinuousLinearEquiv.inverse can be written in terms of Ring.inverse for the ring of self-maps of the domain.

                                                                  theorem ContinuousLinearMap.ring_inverse_eq_map_inverse {R : Type u_3} {M : Type u_4} [TopologicalSpace M] [Ring R] [AddCommGroup M] [Module R M] :
                                                                  Ring.inverse = ContinuousLinearMap.inverse
                                                                  theorem Submodule.ClosedComplemented.exists_submodule_equiv_prod {R : Type u_3} [Ring R] {M : Type u_4} [TopologicalSpace M] [AddCommGroup M] [Module R M] [TopologicalAddGroup M] {p : Submodule R M} (hp : p.ClosedComplemented) :
                                                                  ∃ (q : Submodule R M) (e : M ≃L[R] p × q), (∀ (x : p), e x = (x, 0)) (∀ (y : q), e y = (0, y)) ∀ (x : p × q), e.symm x = x.1 + x.2

                                                                  If p is a closed complemented submodule, then there exists a submodule q and a continuous linear equivalence M ≃L[R] (p × q) such that e (x : p) = (x, 0), e (y : q) = (0, y), and e.symm x = x.1 + x.2.

                                                                  In fact, the properties of e imply the properties of e.symm and vice versa, but we provide both for convenience.

                                                                  The function op is a continuous linear equivalence.

                                                                  Equations
                                                                  Instances For