Documentation

Mathlib.Order.Category.BoolAlg

The category of Boolean algebras #

This defines BoolAlg, the category of Boolean algebras.

structure BoolAlg :
Type (u_1 + 1)

The category of Boolean algebras.

Instances For
    structure BoolAlg.Hom (X Y : BoolAlg) :

    The type of morphisms in BoolAlg R.

    Instances For
      theorem BoolAlg.Hom.ext_iff {X Y : BoolAlg} {x y : X.Hom Y} :
      x = y x.hom' = y.hom'
      theorem BoolAlg.Hom.ext {X Y : BoolAlg} {x y : X.Hom Y} (hom' : x.hom' = y.hom') :
      x = y
      Equations
      • One or more equations did not get rendered due to their size.
      Equations
      • One or more equations did not get rendered due to their size.
      @[reducible, inline]
      abbrev BoolAlg.Hom.hom {X Y : BoolAlg} (f : X.Hom Y) :

      Turn a morphism in BoolAlg back into a BoundedLatticeHom.

      Equations
      Instances For
        @[reducible, inline]
        abbrev BoolAlg.ofHom {X Y : Type u} [BooleanAlgebra X] [BooleanAlgebra Y] (f : BoundedLatticeHom X Y) :
        { carrier := X, str := inst✝ } { carrier := Y, str := inst✝¹ }

        Typecheck a BoundedLatticeHom as a morphism in BoolAlg.

        Equations
        Instances For
          def BoolAlg.Hom.Simps.hom (X Y : BoolAlg) (f : X.Hom Y) :

          Use the ConcreteCategory.hom projection for @[simps] lemmas.

          Equations
          Instances For

            The results below duplicate the ConcreteCategory simp lemmas, but we can keep them for dsimp.

            theorem BoolAlg.ext {X Y : BoolAlg} {f g : X Y} (w : ∀ (x : X), (CategoryTheory.ConcreteCategory.hom f) x = (CategoryTheory.ConcreteCategory.hom g) x) :
            f = g
            theorem BoolAlg.coe_of (X : Type u) [BooleanAlgebra X] :
            { carrier := X, str := inst✝ } = X
            @[simp]
            theorem BoolAlg.hom_comp {X Y Z : BoolAlg} (f : X Y) (g : Y Z) :
            theorem BoolAlg.hom_ext {X Y : BoolAlg} {f g : X Y} (hf : Hom.hom f = Hom.hom g) :
            f = g
            theorem BoolAlg.hom_ext_iff {X Y : BoolAlg} {f g : X Y} :
            @[simp]
            @[simp]
            theorem BoolAlg.ofHom_hom {X Y : BoolAlg} (f : X Y) :
            @[simp]

            Turn a BoolAlg into a BddDistLat by forgetting its complement operation.

            Equations
            Instances For
              Equations
              • One or more equations did not get rendered due to their size.
              Equations
              • One or more equations did not get rendered due to their size.
              @[simp]
              theorem BoolAlg.hasForgetToHeytAlg_forget₂_map {X Y : BoolAlg} (f : X Y) :
              CategoryTheory.HasForget₂.forget₂.map f = HeytAlg.ofHom { toFun := (Hom.hom f), map_sup' := , map_inf' := , map_bot' := , map_himp' := }
              def BoolAlg.Iso.mk {α β : BoolAlg} (e : α ≃o β) :
              α β

              Constructs an equivalence between Boolean algebras from an order isomorphism between them.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                @[simp]
                theorem BoolAlg.Iso.mk_inv {α β : BoolAlg} (e : α ≃o β) :
                (mk e).inv = ofHom (have __src := { toFun := e.symm, map_sup' := , map_inf' := }; { toFun := e.symm, map_sup' := , map_inf' := , map_top' := , map_bot' := })
                @[simp]
                theorem BoolAlg.Iso.mk_hom {α β : BoolAlg} (e : α ≃o β) :
                (mk e).hom = ofHom (have __src := { toFun := e, map_sup' := , map_inf' := }; { toFun := e, map_sup' := , map_inf' := , map_top' := , map_bot' := })

                OrderDual as a functor.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem BoolAlg.dual_map {X✝ Y✝ : BoolAlg} (f : X✝ Y✝) :

                  The equivalence between BoolAlg and itself induced by OrderDual both ways.

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

                    The powerset functor. Set as a contravariant functor.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[simp]
                      theorem typeToBoolAlgOp_map {X Y : Type u} (f : X Y) :
                      typeToBoolAlgOp.map f = (BoolAlg.ofHom (have __src := { toFun := (CompleteLatticeHom.setPreimage f), map_sup' := , map_inf' := }; { toFun := (CompleteLatticeHom.setPreimage f), map_sup' := , map_inf' := , map_top' := , map_bot' := })).op