Documentation

ToMathlib.PFunctor.Lens.Basic

More properties about lenses between polynomial functors #

theorem heq_forall_iff {α : Sort u} {β γ : αSort v} (h : ∀ (a : α), β a = γ a) {f : (a : α) → β a} {g : (a : α) → γ a} :
f g ∀ (a : α), f a g a
theorem PFunctor.Lens.ext {P : PFunctor.{uA₁, uB₁}} {Q : PFunctor.{uA₂, uB₂}} (l₁ l₂ : P.Lens Q) (h₁ : ∀ (a : P.A), l₁.toFunA a = l₂.toFunA a) (h₂ : ∀ (a : P.A), l₁.toFunB a = Eq.rec (motive := fun (x : Q.A) (h : l₂.toFunA a = x) => Q.B xP.B a) (l₂.toFunB a) ) :
l₁ = l₂

The identity lens

Equations
    Instances For

      Composition of lenses

      Equations
        Instances For

          Composition of lenses

          Equations
            Instances For
              structure PFunctor.Lens.Equiv (P : PFunctor.{uA₁, uB₁}) (Q : PFunctor.{uA₂, uB₂}) :
              Type (max (max (max uA₁ uA₂) uB₁) uB₂)

              An equivalence between two polynomial functors P and Q, using lenses. This corresponds to an isomorphism in the category PFunctor with Lens morphisms.

              Instances For
                theorem PFunctor.Lens.Equiv.ext {P : PFunctor.{uA₁, uB₁}} {Q : PFunctor.{uA₂, uB₂}} {x y : P ≃ₗ Q} (toLens : x.toLens = y.toLens) (invLens : x.invLens = y.invLens) :
                x = y

                An equivalence between two polynomial functors P and Q, using lenses. This corresponds to an isomorphism in the category PFunctor with Lens morphisms.

                Equations
                  Instances For
                    Equations
                      Instances For
                        Equations
                          Instances For

                            The (unique) initial lens from the zero functor to any functor P.

                            Equations
                              Instances For

                                The (unique) terminal lens from any functor P to the unit functor 1.

                                Equations
                                  Instances For

                                    Alias of PFunctor.Lens.initial.


                                    The (unique) initial lens from the zero functor to any functor P.

                                    Equations
                                      Instances For

                                        Alias of PFunctor.Lens.terminal.


                                        The (unique) terminal lens from any functor P to the unit functor 1.

                                        Equations
                                          Instances For

                                            Left injection lens inl : P → P + Q

                                            Equations
                                              Instances For

                                                Right injection lens inr : Q → P + Q

                                                Equations
                                                  Instances For

                                                    Copairing of lenses [l₁, l₂]ₗ : P + Q → R

                                                    Equations
                                                      Instances For

                                                        Parallel application of lenses for coproduct l₁ ⊎ l₂ : P + Q → R + W

                                                        Equations
                                                          Instances For

                                                            Projection lens fst : P * Q → P

                                                            Equations
                                                              Instances For

                                                                Projection lens snd : P * Q → Q

                                                                Equations
                                                                  Instances For

                                                                    Pairing of lenses ⟨l₁, l₂⟩ₗ : P → Q * R

                                                                    Equations
                                                                      Instances For

                                                                        Parallel application of lenses for product l₁ ×ₗ l₂ : P * Q → R * W

                                                                        Equations
                                                                          Instances For

                                                                            Apply lenses to both sides of a composition: l₁ ◃ₗ l₂ : (P ◃ Q ⇆ R ◃ W)

                                                                            Equations
                                                                              Instances For

                                                                                Apply lenses to both sides of a tensor / parallel product: l₁ ⊗ₗ l₂ : (P ⊗ Q ⇆ R ⊗ W)

                                                                                Equations
                                                                                  Instances For

                                                                                    Lens to introduce X on the right: P → P ◃ X

                                                                                    Equations
                                                                                      Instances For

                                                                                        Lens to introduce X on the left: P → X ◃ P

                                                                                        Equations
                                                                                          Instances For

                                                                                            Lens from P ◃ X to P

                                                                                            Equations
                                                                                              Instances For

                                                                                                Lens from X ◃ P to P

                                                                                                Equations
                                                                                                  Instances For

                                                                                                    Apply lenses to both sides of a composition: l₁ ◃ₗ l₂ : (P ◃ Q ⇆ R ◃ W)

                                                                                                    Equations
                                                                                                      Instances For

                                                                                                        Parallel application of lenses for product l₁ ×ₗ l₂ : P * Q → R * W

                                                                                                        Equations
                                                                                                          Instances For

                                                                                                            Parallel application of lenses for coproduct l₁ ⊎ l₂ : P + Q → R + W

                                                                                                            Equations
                                                                                                              Instances For

                                                                                                                Apply lenses to both sides of a tensor / parallel product: l₁ ⊗ₗ l₂ : (P ⊗ Q ⇆ R ⊗ W)

                                                                                                                Equations
                                                                                                                  Instances For
                                                                                                                    def PFunctor.Lens.enclose (P : PFunctor.{uA, uB}) :
                                                                                                                    Type (max uA uA₁ uB uB₁)

                                                                                                                    The type of lenses from a polynomial functor P to X

                                                                                                                    Equations
                                                                                                                      Instances For

                                                                                                                        Helper lens for speedup

                                                                                                                        Equations
                                                                                                                          Instances For

                                                                                                                            The speedup lens operation: Lens (S y^S) P → Lens (S y^S) (P ◃ P)

                                                                                                                            Equations
                                                                                                                              Instances For

                                                                                                                                Commutativity of coproduct

                                                                                                                                Equations
                                                                                                                                  Instances For

                                                                                                                                    Associativity of coproduct

                                                                                                                                    Equations
                                                                                                                                      Instances For

                                                                                                                                        Coproduct with 0 is identity (right)

                                                                                                                                        Equations
                                                                                                                                          Instances For

                                                                                                                                            Coproduct with 0 is identity (left)

                                                                                                                                            Equations
                                                                                                                                              Instances For

                                                                                                                                                Commutativity of product

                                                                                                                                                Equations
                                                                                                                                                  Instances For

                                                                                                                                                    Associativity of product

                                                                                                                                                    Equations
                                                                                                                                                      Instances For

                                                                                                                                                        Product with 1 is identity (right)

                                                                                                                                                        Equations
                                                                                                                                                          Instances For

                                                                                                                                                            Product with 1 is identity (left)

                                                                                                                                                            Equations
                                                                                                                                                              Instances For

                                                                                                                                                                Product with 0 is zero (right)

                                                                                                                                                                Equations
                                                                                                                                                                  Instances For

                                                                                                                                                                    Product with 0 is zero (left)

                                                                                                                                                                    Equations
                                                                                                                                                                      Instances For

                                                                                                                                                                        Left distributive law for product over coproduct

                                                                                                                                                                        Equations
                                                                                                                                                                          Instances For

                                                                                                                                                                            Right distributive law for coproduct over product

                                                                                                                                                                            Equations
                                                                                                                                                                              Instances For
                                                                                                                                                                                theorem PFunctor.Lens.compMap_comp {P : PFunctor.{uA₁, uB₁}} {Q : PFunctor.{uA₂, uB₂}} {R : PFunctor.{uA₃, uB₃}} {P' : PFunctor.{uA₄, uB₄}} {Q' : PFunctor.{uA₅, uB₅}} {R' : PFunctor.{uA₆, uB₆}} (l₁ : P.Lens P') (l₂ : Q.Lens Q') (l₁' : P'.Lens R) (l₂' : Q'.Lens R') :
                                                                                                                                                                                l₁' ∘ₗ l₁ ◃ₗ (l₂' ∘ₗ l₂) = l₁' ◃ₗ l₂' ∘ₗ (l₁ ◃ₗ l₂)

                                                                                                                                                                                Associativity of composition

                                                                                                                                                                                Equations
                                                                                                                                                                                  Instances For

                                                                                                                                                                                    Composition with X is identity (right)

                                                                                                                                                                                    Equations
                                                                                                                                                                                      Instances For

                                                                                                                                                                                        Composition with X is identity (left)

                                                                                                                                                                                        Equations
                                                                                                                                                                                          Instances For

                                                                                                                                                                                            Distributivity of composition over coproduct on the right

                                                                                                                                                                                            Equations
                                                                                                                                                                                              Instances For
                                                                                                                                                                                                theorem PFunctor.Lens.tensorMap_comp {P : PFunctor.{uA₁, uB₁}} {Q : PFunctor.{uA₂, uB₂}} {R : PFunctor.{uA₃, uB₃}} {P' : PFunctor.{uA₄, uB₄}} {Q' : PFunctor.{uA₅, uB₅}} {R' : PFunctor.{uA₆, uB₆}} (l₁ : P.Lens P') (l₂ : Q.Lens Q') (l₁' : P'.Lens R) (l₂' : Q'.Lens R') :
                                                                                                                                                                                                l₁' ∘ₗ l₁ ⊗ₗ (l₂' ∘ₗ l₂) = l₁' ⊗ₗ l₂' ∘ₗ (l₁ ⊗ₗ l₂)

                                                                                                                                                                                                Commutativity of tensor product

                                                                                                                                                                                                Equations
                                                                                                                                                                                                  Instances For

                                                                                                                                                                                                    Associativity of tensor product

                                                                                                                                                                                                    Equations
                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                        Tensor product with X is identity (right)

                                                                                                                                                                                                        Equations
                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                            Tensor product with X is identity (left)

                                                                                                                                                                                                            Equations
                                                                                                                                                                                                              Instances For

                                                                                                                                                                                                                Tensor product with 0 is zero (left)

                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                  Instances For

                                                                                                                                                                                                                    Tensor product with 0 is zero (right)

                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                                        Left distributivity of tensor product over coproduct

                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                            Right distributivity of tensor product over coproduct

                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                instance PFunctor.Lens.instIsEmptyBSigma {I : Type v} [IsEmpty I] {F : IPFunctor.{u, u_1}} {a : (sigma F).A} :
                                                                                                                                                                                                                                IsEmpty ((sigma F).B a)

                                                                                                                                                                                                                                Pi over a PUnit-indexed family is equivalent to the functor itself.

                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                  Instances For

                                                                                                                                                                                                                                    ULift equivalence for lenses

                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                                                        Convert an equivalence between two polynomial functors P and Q to a lens.

                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                          Instances For