Documentation

Mathlib.Order.BooleanSubalgebra

Boolean subalgebras #

This file defines boolean subalgebras.

structure BooleanSubalgebra (α : Type u_2) [BooleanAlgebra α] extends Sublattice α :
Type u_2

A boolean subalgebra of a boolean algebra is a set containing the bottom and top elements, and closed under suprema, infima and complements.

Instances For
    Equations
    • BooleanSubalgebra.instSetLike = { coe := fun (L : BooleanSubalgebra α) => L.carrier, coe_injective' := }
    theorem BooleanSubalgebra.coe_inj {α : Type u_2} [BooleanAlgebra α] {L M : BooleanSubalgebra α} :
    L = M L = M
    theorem BooleanSubalgebra.compl_mem {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} {a : α} (ha : a L) :
    a L
    @[simp]
    theorem BooleanSubalgebra.compl_mem_iff {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} {a : α} :
    a L a L
    @[simp]
    @[simp]
    theorem BooleanSubalgebra.sup_mem {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} {a b : α} (ha : a L) (hb : b L) :
    a b L
    theorem BooleanSubalgebra.inf_mem {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} {a b : α} (ha : a L) (hb : b L) :
    a b L
    theorem BooleanSubalgebra.sdiff_mem {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} {a b : α} (ha : a L) (hb : b L) :
    a \ b L
    theorem BooleanSubalgebra.himp_mem {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} {a b : α} (ha : a L) (hb : b L) :
    a b L
    theorem BooleanSubalgebra.mem_carrier {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} {a : α} :
    a L.carrier a L
    @[simp]
    theorem BooleanSubalgebra.mem_toSublattice {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} {a : α} :
    a L.toSublattice a L
    @[simp]
    theorem BooleanSubalgebra.mem_mk {α : Type u_2} [BooleanAlgebra α] {a : α} {L : Sublattice α} (h_compl : ∀ {a : α}, a L.carriera L.carrier) (h_bot : L.carrier) :
    a { toSublattice := L, compl_mem' := h_compl, bot_mem' := h_bot } a L
    @[simp]
    theorem BooleanSubalgebra.coe_mk {α : Type u_2} [BooleanAlgebra α] (L : Sublattice α) (h_compl : ∀ {a : α}, a L.carriera L.carrier) (h_bot : L.carrier) :
    { toSublattice := L, compl_mem' := h_compl, bot_mem' := h_bot } = L
    @[simp]
    theorem BooleanSubalgebra.mk_le_mk {α : Type u_2} [BooleanAlgebra α] {L M : Sublattice α} (hL_compl : ∀ {a : α}, a L.carriera L.carrier) (hL_bot : L.carrier) (hM_compl : ∀ {a : α}, a M.carriera M.carrier) (hM_bot : M.carrier) :
    { toSublattice := L, compl_mem' := hL_compl, bot_mem' := hL_bot } { toSublattice := M, compl_mem' := hM_compl, bot_mem' := hM_bot } L M
    @[simp]
    theorem BooleanSubalgebra.mk_lt_mk {α : Type u_2} [BooleanAlgebra α] {L M : Sublattice α} (hL_compl : ∀ {a : α}, a L.carriera L.carrier) (hL_bot : L.carrier) (hM_compl : ∀ {a : α}, a M.carriera M.carrier) (hM_bot : M.carrier) :
    { toSublattice := L, compl_mem' := hL_compl, bot_mem' := hL_bot } < { toSublattice := M, compl_mem' := hM_compl, bot_mem' := hM_bot } L < M
    def BooleanSubalgebra.copy {α : Type u_2} [BooleanAlgebra α] (L : BooleanSubalgebra α) (s : Set α) (hs : s = L) :

    Copy of a boolean subalgebra with a new carrier equal to the old one. Useful to fix definitional equalities.

    Equations
    • L.copy s hs = { toSublattice := L.copy s , compl_mem' := , bot_mem' := }
    Instances For
      @[simp]
      theorem BooleanSubalgebra.coe_copy {α : Type u_2} [BooleanAlgebra α] (L : BooleanSubalgebra α) (s : Set α) (hs : s = L) :
      (L.copy s hs) = s
      theorem BooleanSubalgebra.copy_eq {α : Type u_2} [BooleanAlgebra α] (L : BooleanSubalgebra α) (s : Set α) (hs : s = L) :
      L.copy s hs = L
      theorem BooleanSubalgebra.ext {α : Type u_2} [BooleanAlgebra α] {L M : BooleanSubalgebra α} :
      (∀ (a : α), a L a M)L = M

      Two boolean subalgebras are equal if they have the same elements.

      instance BooleanSubalgebra.instBotCoe {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} :
      Bot L

      A boolean subalgebra of a lattice inherits a bottom element.

      Equations
      • BooleanSubalgebra.instBotCoe = { bot := , }
      instance BooleanSubalgebra.instTopCoe {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} :
      Top L

      A boolean subalgebra of a lattice inherits a top element.

      Equations
      • BooleanSubalgebra.instTopCoe = { top := , }
      instance BooleanSubalgebra.instSupCoe {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} :
      Max L

      A boolean subalgebra of a lattice inherits a supremum.

      Equations
      • BooleanSubalgebra.instSupCoe = { max := fun (a b : L) => a b, }
      instance BooleanSubalgebra.instInfCoe {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} :
      Min L

      A boolean subalgebra of a lattice inherits an infimum.

      Equations
      • BooleanSubalgebra.instInfCoe = { min := fun (a b : L) => a b, }

      A boolean subalgebra of a lattice inherits a complement.

      Equations
      • BooleanSubalgebra.instHasComplCoe = { compl := fun (a : L) => (↑a), }

      A boolean subalgebra of a lattice inherits a difference.

      Equations
      • BooleanSubalgebra.instSDiffCoe = { sdiff := fun (a b : L) => a \ b, }

      A boolean subalgebra of a lattice inherits a Heyting implication.

      Equations
      • BooleanSubalgebra.instHImpCoe = { himp := fun (a b : L) => a b, }
      @[simp]
      @[simp]
      @[simp]
      theorem BooleanSubalgebra.val_sup {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a b : L) :
      (a b) = a b
      @[simp]
      theorem BooleanSubalgebra.val_inf {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a b : L) :
      (a b) = a b
      @[simp]
      theorem BooleanSubalgebra.val_compl {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a : L) :
      a = (↑a)
      @[simp]
      theorem BooleanSubalgebra.val_sdiff {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a b : L) :
      (a \ b) = a \ b
      @[simp]
      theorem BooleanSubalgebra.val_himp {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a b : L) :
      (a b) = a b
      @[simp]
      theorem BooleanSubalgebra.mk_bot {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} :
      , =
      @[simp]
      theorem BooleanSubalgebra.mk_top {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} :
      , =
      @[simp]
      theorem BooleanSubalgebra.mk_sup_mk {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a b : α) (ha : a L) (hb : b L) :
      a, ha b, hb = a b,
      @[simp]
      theorem BooleanSubalgebra.mk_inf_mk {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a b : α) (ha : a L) (hb : b L) :
      a, ha b, hb = a b,
      @[simp]
      theorem BooleanSubalgebra.compl_mk {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a : α) (ha : a L) :
      a, ha = a,
      @[simp]
      theorem BooleanSubalgebra.mk_sdiff_mk {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a b : α) (ha : a L) (hb : b L) :
      a, ha \ b, hb = a \ b,
      @[simp]
      theorem BooleanSubalgebra.mk_himp_mk {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} (a b : α) (ha : a L) (hb : b L) :
      a, ha b, hb = a b,

      A boolean subalgebra of a lattice inherits a boolean algebra structure.

      Equations

      The natural lattice hom from a boolean subalgebra to the original lattice.

      Equations
      • L.subtype = { toFun := Subtype.val, map_sup' := , map_inf' := , map_top' := , map_bot' := }
      Instances For
        @[simp]
        theorem BooleanSubalgebra.coe_subtype {α : Type u_2} [BooleanAlgebra α] (L : BooleanSubalgebra α) :
        L.subtype = Subtype.val
        theorem BooleanSubalgebra.subtype_apply {α : Type u_2} [BooleanAlgebra α] (L : BooleanSubalgebra α) (a : L) :
        L.subtype a = a
        def BooleanSubalgebra.inclusion {α : Type u_2} [BooleanAlgebra α] {L M : BooleanSubalgebra α} (h : L M) :

        The inclusion homomorphism from a boolean subalgebra L to a bigger boolean subalgebra M.

        Equations
        Instances For
          @[simp]
          theorem BooleanSubalgebra.subtype_comp_inclusion {α : Type u_2} [BooleanAlgebra α] {L M : BooleanSubalgebra α} (h : L M) :
          M.subtype.comp (BooleanSubalgebra.inclusion h) = L.subtype

          The maximum boolean subalgebra of a lattice.

          Equations
          • BooleanSubalgebra.instTop = { top := { carrier := Set.univ, supClosed' := , infClosed' := , compl_mem' := , bot_mem' := } }

          The trivial boolean subalgebra of a lattice.

          Equations
          • BooleanSubalgebra.instBot = { bot := { carrier := {, }, supClosed' := , infClosed' := , compl_mem' := , bot_mem' := } }

          The inf of two boolean subalgebras is their intersection.

          Equations
          • BooleanSubalgebra.instInf = { min := fun (L M : BooleanSubalgebra α) => { carrier := L M, supClosed' := , infClosed' := , compl_mem' := , bot_mem' := } }

          The inf of boolean subalgebras is their intersection.

          Equations
          • BooleanSubalgebra.instInfSet = { sInf := fun (S : Set (BooleanSubalgebra α)) => { carrier := LS, L, supClosed' := , infClosed' := , compl_mem' := , bot_mem' := } }
          Equations
          • BooleanSubalgebra.instInhabited = { default := }

          The top boolean subalgebra is isomorphic to the original boolean algebra.

          This is the boolean subalgebra version of Equiv.Set.univ α.

          Equations
          • BooleanSubalgebra.topEquiv = { toEquiv := Equiv.Set.univ α, map_rel_iff' := }
          Instances For
            @[simp]
            theorem BooleanSubalgebra.coe_top {α : Type u_2} [BooleanAlgebra α] :
            = Set.univ
            @[simp]
            theorem BooleanSubalgebra.coe_bot {α : Type u_2} [BooleanAlgebra α] :
            = {, }
            @[simp]
            theorem BooleanSubalgebra.coe_inf {α : Type u_2} [BooleanAlgebra α] (L M : BooleanSubalgebra α) :
            (L M) = L M
            @[simp]
            theorem BooleanSubalgebra.coe_sInf {α : Type u_2} [BooleanAlgebra α] (S : Set (BooleanSubalgebra α)) :
            (sInf S) = LS, L
            @[simp]
            theorem BooleanSubalgebra.coe_iInf {ι : Sort u_1} {α : Type u_2} [BooleanAlgebra α] (f : ιBooleanSubalgebra α) :
            (⨅ (i : ι), f i) = ⋂ (i : ι), (f i)
            @[simp]
            theorem BooleanSubalgebra.coe_eq_univ {α : Type u_2} [BooleanAlgebra α] {L : BooleanSubalgebra α} :
            L = Set.univ L =
            @[simp]
            theorem BooleanSubalgebra.mem_bot {α : Type u_2} [BooleanAlgebra α] {a : α} :
            @[simp]
            theorem BooleanSubalgebra.mem_top {α : Type u_2} [BooleanAlgebra α] {a : α} :
            @[simp]
            theorem BooleanSubalgebra.mem_inf {α : Type u_2} [BooleanAlgebra α] {L M : BooleanSubalgebra α} {a : α} :
            a L M a L a M
            @[simp]
            theorem BooleanSubalgebra.mem_sInf {α : Type u_2} [BooleanAlgebra α] {a : α} {S : Set (BooleanSubalgebra α)} :
            a sInf S LS, a L
            @[simp]
            theorem BooleanSubalgebra.mem_iInf {ι : Sort u_1} {α : Type u_2} [BooleanAlgebra α] {a : α} {f : ιBooleanSubalgebra α} :
            a ⨅ (i : ι), f i ∀ (i : ι), a f i

            BooleanSubalgebras of a lattice form a complete lattice.

            Equations
            Equations

            The preimage of a boolean subalgebra along a bounded lattice homomorphism.

            Equations
            Instances For
              @[simp]
              theorem BooleanSubalgebra.coe_comap {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] (L : BooleanSubalgebra β) (f : BoundedLatticeHom α β) :
              (BooleanSubalgebra.comap f L) = f ⁻¹' L
              @[simp]
              theorem BooleanSubalgebra.mem_comap {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] {f : BoundedLatticeHom α β} {a : α} {L : BooleanSubalgebra β} :

              The image of a boolean subalgebra along a monoid homomorphism is a boolean subalgebra.

              Equations
              • BooleanSubalgebra.map f L = { carrier := f '' L, supClosed' := , infClosed' := , compl_mem' := , bot_mem' := }
              Instances For
                @[simp]
                theorem BooleanSubalgebra.coe_map {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] (f : BoundedLatticeHom α β) (L : BooleanSubalgebra α) :
                (BooleanSubalgebra.map f L) = f '' L
                @[simp]
                theorem BooleanSubalgebra.mem_map {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] {L : BooleanSubalgebra α} {f : BoundedLatticeHom α β} {b : β} :
                b BooleanSubalgebra.map f L aL, f a = b
                theorem BooleanSubalgebra.mem_map_of_mem {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] {L : BooleanSubalgebra α} (f : BoundedLatticeHom α β) {a : α} :
                a Lf a BooleanSubalgebra.map f L
                theorem BooleanSubalgebra.apply_coe_mem_map {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] {L : BooleanSubalgebra α} (f : BoundedLatticeHom α β) (a : L) :
                theorem BooleanSubalgebra.mem_map_equiv {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] {L : BooleanSubalgebra α} {f : α ≃o β} {a : β} :
                a BooleanSubalgebra.map (let __src := { toFun := f, map_sup' := , map_inf' := }; { toFun := f, map_sup' := , map_inf' := , map_top' := , map_bot' := }) L f.symm a L
                theorem BooleanSubalgebra.apply_mem_map_iff {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] {L : BooleanSubalgebra α} {f : BoundedLatticeHom α β} {a : α} (hf : Function.Injective f) :
                theorem BooleanSubalgebra.map_equiv_eq_comap_symm {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] (f : α ≃o β) (L : BooleanSubalgebra α) :
                BooleanSubalgebra.map (let __src := { toFun := f, map_sup' := , map_inf' := }; { toFun := f, map_sup' := , map_inf' := , map_top' := , map_bot' := }) L = BooleanSubalgebra.comap (let __src := { toFun := f.symm, map_sup' := , map_inf' := }; { toFun := f.symm, map_sup' := , map_inf' := , map_top' := , map_bot' := }) L
                theorem BooleanSubalgebra.comap_equiv_eq_map_symm {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] (f : β ≃o α) (L : BooleanSubalgebra α) :
                BooleanSubalgebra.comap (let __src := { toFun := f, map_sup' := , map_inf' := }; { toFun := f, map_sup' := , map_inf' := , map_top' := , map_bot' := }) L = BooleanSubalgebra.map (let __src := { toFun := f.symm, map_sup' := , map_inf' := }; { toFun := f.symm, map_sup' := , map_inf' := , map_top' := , map_bot' := }) L
                theorem BooleanSubalgebra.map_symm_eq_iff_eq_map {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] {L : BooleanSubalgebra α} {M : BooleanSubalgebra β} {e : β ≃o α} :
                BooleanSubalgebra.map (let __src := { toFun := e.symm, map_sup' := , map_inf' := }; { toFun := e.symm, map_sup' := , map_inf' := , map_top' := , map_bot' := }) L = M L = BooleanSubalgebra.map (let __src := { toFun := e, map_sup' := , map_inf' := }; { toFun := e, map_sup' := , map_inf' := , map_top' := , map_bot' := }) M
                theorem BooleanSubalgebra.map_iSup {ι : Sort u_1} {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] (f : BoundedLatticeHom α β) (L : ιBooleanSubalgebra α) :
                BooleanSubalgebra.map f (⨆ (i : ι), L i) = ⨆ (i : ι), BooleanSubalgebra.map f (L i)
                theorem BooleanSubalgebra.comap_iInf {ι : Sort u_1} {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] (f : BoundedLatticeHom α β) (L : ιBooleanSubalgebra β) :
                BooleanSubalgebra.comap f (⨅ (i : ι), L i) = ⨅ (i : ι), BooleanSubalgebra.comap f (L i)
                theorem BooleanSubalgebra.le_comap_iSup {ι : Sort u_1} {α : Type u_2} {β : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] (f : BoundedLatticeHom α β) (L : ιBooleanSubalgebra β) :
                ⨆ (i : ι), BooleanSubalgebra.comap f (L i) BooleanSubalgebra.comap f (⨆ (i : ι), L i)

                The minimum boolean subalgebra containing a given set.

                Equations
                Instances For
                  theorem BooleanSubalgebra.mem_closure {α : Type u_2} [BooleanAlgebra α] {s : Set α} {x : α} :
                  x BooleanSubalgebra.closure s ∀ ⦃L : BooleanSubalgebra α⦄, s Lx L
                  theorem BooleanSubalgebra.closure_bot_sup_induction {α : Type u_2} [BooleanAlgebra α] {s : Set α} {p : (g : α) → g BooleanSubalgebra.closure sProp} (mem : ∀ (x : α) (hx : x s), p x ) (bot : p ) (sup : ∀ (x : α) (hx : x BooleanSubalgebra.closure s) (y : α) (hy : y BooleanSubalgebra.closure s), p x hxp y hyp (x y) ) (compl : ∀ (x : α) (hx : x BooleanSubalgebra.closure s), p x hxp x ) {x : α} (hx : x BooleanSubalgebra.closure s) :
                  p x hx

                  An induction principle for closure membership. If p holds for and all elements of s, and is preserved under suprema and complement, then p holds for all elements of the closure of s.

                  theorem BooleanSubalgebra.iSup_mem {ι : Sort u_1} {α : Type u_2} [CompleteBooleanAlgebra α] {L : BooleanSubalgebra α} {f : ια} [Finite ι] (hf : ∀ (i : ι), f i L) :
                  ⨆ (i : ι), f i L
                  theorem BooleanSubalgebra.iInf_mem {ι : Sort u_1} {α : Type u_2} [CompleteBooleanAlgebra α] {L : BooleanSubalgebra α} {f : ια} [Finite ι] (hf : ∀ (i : ι), f i L) :
                  ⨅ (i : ι), f i L
                  theorem BooleanSubalgebra.sSup_mem {α : Type u_2} [CompleteBooleanAlgebra α] {L : BooleanSubalgebra α} {s : Set α} (hs : s.Finite) (hsL : s L) :
                  sSup s L
                  theorem BooleanSubalgebra.sInf_mem {α : Type u_2} [CompleteBooleanAlgebra α] {L : BooleanSubalgebra α} {s : Set α} (hs : s.Finite) (hsL : s L) :
                  sInf s L
                  theorem BooleanSubalgebra.biSup_mem {α : Type u_2} [CompleteBooleanAlgebra α] {L : BooleanSubalgebra α} {ι : Type u_5} {t : Set ι} {f : ια} (ht : t.Finite) (hf : it, f i L) :
                  it, f i L
                  theorem BooleanSubalgebra.biInf_mem {α : Type u_2} [CompleteBooleanAlgebra α] {L : BooleanSubalgebra α} {ι : Type u_5} {t : Set ι} {f : ια} (ht : t.Finite) (hf : it, f i L) :
                  it, f i L