Documentation

Mathlib.Order.Category.HeytAlg

The category of Heyting algebras #

This file defines HeytAlg, the category of Heyting algebras.

structure HeytAlg :
Type (u_1 + 1)

The category of Heyting algebras.

Instances For
    @[reducible, inline]
    abbrev HeytAlg.of (X : Type u_1) [HeytingAlgebra X] :

    Construct a bundled HeytAlg from the underlying type and typeclass.

    Equations
    Instances For
      structure HeytAlg.Hom (X Y : HeytAlg) :

      The type of morphisms in HeytAlg R.

      Instances For
        theorem HeytAlg.Hom.ext {X Y : HeytAlg} {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.
        @[reducible, inline]
        abbrev HeytAlg.Hom.hom {X Y : HeytAlg} (f : X.Hom Y) :
        HeytingHom X Y

        Turn a morphism in HeytAlg back into a HeytingHom.

        Equations
        Instances For
          @[reducible, inline]
          abbrev HeytAlg.ofHom {X Y : Type u} [HeytingAlgebra X] [HeytingAlgebra Y] (f : HeytingHom X Y) :
          of X of Y

          Typecheck a HeytingHom as a morphism in HeytAlg.

          Equations
          Instances For
            def HeytAlg.Hom.Simps.hom (X Y : HeytAlg) (f : X.Hom Y) :
            HeytingHom X 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 HeytAlg.ext {X Y : HeytAlg} {f g : X Y} (w : ∀ (x : X), (CategoryTheory.ConcreteCategory.hom f) x = (CategoryTheory.ConcreteCategory.hom g) x) :
              f = g
              theorem HeytAlg.coe_of (X : Type u) [HeytingAlgebra X] :
              (of X) = X
              @[simp]
              theorem HeytAlg.hom_comp {X Y Z : HeytAlg} (f : X Y) (g : Y Z) :
              theorem HeytAlg.hom_ext {X Y : HeytAlg} {f g : X Y} (hf : Hom.hom f = Hom.hom g) :
              f = g
              @[simp]
              theorem HeytAlg.hom_ofHom {X Y : Type u} [HeytingAlgebra X] [HeytingAlgebra Y] (f : HeytingHom X Y) :
              @[simp]
              theorem HeytAlg.ofHom_hom {X Y : HeytAlg} (f : X Y) :
              Equations
              • One or more equations did not get rendered due to their size.
              @[simp]
              theorem HeytAlg.hasForgetToLat_forget₂_map {X✝ Y✝ : HeytAlg} (f : X✝ Y✝) :
              CategoryTheory.HasForget₂.forget₂.map f = BddDistLat.ofHom (let __src := { toFun := (Hom.hom f), map_sup' := , map_inf' := }; { toFun := (Hom.hom f), map_sup' := , map_inf' := , map_top' := , map_bot' := })
              def HeytAlg.Iso.mk {α β : HeytAlg} (e : α ≃o β) :
              α β

              Constructs an isomorphism of Heyting algebras from an order isomorphism between them.

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