Documentation

Mathlib.Algebra.Ring.BooleanRing

Boolean rings #

A Boolean ring is a ring where multiplication is idempotent. They are equivalent to Boolean algebras.

Main declarations #

Implementation notes #

We provide two ways of turning a Boolean algebra/ring into a Boolean ring/algebra:

At this point in time, it is not clear the first way is useful, but we keep it for educational purposes and because it is easier than dealing with ofBoolAlg/toBoolAlg/ofBoolRing/toBoolRing explicitly.

Tags #

boolean ring, boolean algebra

theorem BooleanRing.mul_self {α : Type u_4} [self : BooleanRing α] (a : α) :
a * a = a

Multiplication in a boolean ring is idempotent.

instance BooleanRing.instIdempotentOpHMul {α : Type u_1} [BooleanRing α] :
Std.IdempotentOp fun (x1 x2 : α) => x1 * x2
Equations
  • =
theorem BooleanRing.add_self {α : Type u_1} [BooleanRing α] (a : α) :
a + a = 0
theorem BooleanRing.neg_eq {α : Type u_1} [BooleanRing α] (a : α) :
-a = a
theorem BooleanRing.add_eq_zero' {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
a + b = 0 a = b
@[simp]
theorem BooleanRing.mul_add_mul {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
a * b + b * a = 0
theorem BooleanRing.sub_eq_add {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
a - b = a + b
@[simp]
theorem BooleanRing.mul_one_add_self {α : Type u_1} [BooleanRing α] (a : α) :
a * (1 + a) = 0
@[instance 100]
instance BooleanRing.toCommRing {α : Type u_1} [BooleanRing α] :
Equations

Turning a Boolean ring into a Boolean algebra #

def AsBoolAlg (α : Type u_4) :
Type u_4

Type synonym to view a Boolean ring as a Boolean algebra.

Equations
Instances For
    def toBoolAlg {α : Type u_1} :

    The "identity" equivalence between AsBoolAlg α and α.

    Equations
    Instances For
      def ofBoolAlg {α : Type u_1} :

      The "identity" equivalence between α and AsBoolAlg α.

      Equations
      Instances For
        @[simp]
        theorem toBoolAlg_symm_eq {α : Type u_1} :
        toBoolAlg.symm = ofBoolAlg
        @[simp]
        theorem ofBoolAlg_symm_eq {α : Type u_1} :
        ofBoolAlg.symm = toBoolAlg
        @[simp]
        theorem toBoolAlg_ofBoolAlg {α : Type u_1} (a : AsBoolAlg α) :
        toBoolAlg (ofBoolAlg a) = a
        @[simp]
        theorem ofBoolAlg_toBoolAlg {α : Type u_1} (a : α) :
        ofBoolAlg (toBoolAlg a) = a
        theorem toBoolAlg_inj {α : Type u_1} {a : α} {b : α} :
        toBoolAlg a = toBoolAlg b a = b
        theorem ofBoolAlg_inj {α : Type u_1} {a : AsBoolAlg α} {b : AsBoolAlg α} :
        ofBoolAlg a = ofBoolAlg b a = b
        Equations
        • instInhabitedAsBoolAlg = inst
        def BooleanRing.sup {α : Type u_1} [BooleanRing α] :
        Sup α

        The join operation in a Boolean ring is x + y + x * y.

        Equations
        • BooleanRing.sup = { sup := fun (x y : α) => x + y + x * y }
        Instances For
          def BooleanRing.inf {α : Type u_1} [BooleanRing α] :
          Inf α

          The meet operation in a Boolean ring is x * y.

          Equations
          • BooleanRing.inf = { inf := fun (x1 x2 : α) => x1 * x2 }
          Instances For
            theorem BooleanRing.sup_comm {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
            a b = b a
            theorem BooleanRing.inf_comm {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
            a b = b a
            theorem BooleanRing.sup_assoc {α : Type u_1} [BooleanRing α] (a : α) (b : α) (c : α) :
            a b c = a (b c)
            theorem BooleanRing.inf_assoc {α : Type u_1} [BooleanRing α] (a : α) (b : α) (c : α) :
            a b c = a (b c)
            theorem BooleanRing.sup_inf_self {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
            a a b = a
            theorem BooleanRing.inf_sup_self {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
            a (a b) = a
            theorem BooleanRing.le_sup_inf_aux {α : Type u_1} [BooleanRing α] (a : α) (b : α) (c : α) :
            (a + b + a * b) * (a + c + a * c) = a + b * c + a * (b * c)
            theorem BooleanRing.le_sup_inf {α : Type u_1} [BooleanRing α] (a : α) (b : α) (c : α) :
            (a b) (a c) (a b c) = a b c

            The Boolean algebra structure on a Boolean ring.

            The data is defined so that:

            • a ⊔ b unfolds to a + b + a * b
            • a ⊓ b unfolds to a * b
            • a ≤ b unfolds to a + b + a * b = b
            • unfolds to 0
            • unfolds to 1
            • aᶜ unfolds to 1 + a
            • a \ b unfolds to a * (1 + b)
            Equations
            Instances For
              Equations
              • instBooleanAlgebraAsBoolAlg = BooleanRing.toBooleanAlgebra
              @[simp]
              theorem ofBoolAlg_top {α : Type u_1} [BooleanRing α] :
              ofBoolAlg = 1
              @[simp]
              theorem ofBoolAlg_bot {α : Type u_1} [BooleanRing α] :
              ofBoolAlg = 0
              @[simp]
              theorem ofBoolAlg_sup {α : Type u_1} [BooleanRing α] (a : AsBoolAlg α) (b : AsBoolAlg α) :
              ofBoolAlg (a b) = ofBoolAlg a + ofBoolAlg b + ofBoolAlg a * ofBoolAlg b
              @[simp]
              theorem ofBoolAlg_inf {α : Type u_1} [BooleanRing α] (a : AsBoolAlg α) (b : AsBoolAlg α) :
              ofBoolAlg (a b) = ofBoolAlg a * ofBoolAlg b
              @[simp]
              theorem ofBoolAlg_compl {α : Type u_1} [BooleanRing α] (a : AsBoolAlg α) :
              ofBoolAlg a = 1 + ofBoolAlg a
              @[simp]
              theorem ofBoolAlg_sdiff {α : Type u_1} [BooleanRing α] (a : AsBoolAlg α) (b : AsBoolAlg α) :
              ofBoolAlg (a \ b) = ofBoolAlg a * (1 + ofBoolAlg b)
              @[simp]
              theorem ofBoolAlg_symmDiff {α : Type u_1} [BooleanRing α] (a : AsBoolAlg α) (b : AsBoolAlg α) :
              ofBoolAlg (symmDiff a b) = ofBoolAlg a + ofBoolAlg b
              @[simp]
              theorem ofBoolAlg_mul_ofBoolAlg_eq_left_iff {α : Type u_1} [BooleanRing α] {a : AsBoolAlg α} {b : AsBoolAlg α} :
              ofBoolAlg a * ofBoolAlg b = ofBoolAlg a a b
              @[simp]
              theorem toBoolAlg_zero {α : Type u_1} [BooleanRing α] :
              toBoolAlg 0 =
              @[simp]
              theorem toBoolAlg_one {α : Type u_1} [BooleanRing α] :
              toBoolAlg 1 =
              @[simp]
              theorem toBoolAlg_mul {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
              toBoolAlg (a * b) = toBoolAlg a toBoolAlg b
              @[simp]
              theorem toBoolAlg_add_add_mul {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
              toBoolAlg (a + b + a * b) = toBoolAlg a toBoolAlg b
              @[simp]
              theorem toBoolAlg_add {α : Type u_1} [BooleanRing α] (a : α) (b : α) :
              toBoolAlg (a + b) = symmDiff (toBoolAlg a) (toBoolAlg b)
              def RingHom.asBoolAlg {α : Type u_1} {β : Type u_2} [BooleanRing α] [BooleanRing β] (f : α →+* β) :

              Turn a ring homomorphism from Boolean rings α to β into a bounded lattice homomorphism from α to β considered as Boolean algebras.

              Equations
              • f.asBoolAlg = { toFun := toBoolAlg f ofBoolAlg, map_sup' := , map_inf' := , map_top' := , map_bot' := }
              Instances For
                @[simp]
                theorem RingHom.asBoolAlg_toLatticeHom_toSupHom_toFun {α : Type u_1} {β : Type u_2} [BooleanRing α] [BooleanRing β] (f : α →+* β) :
                ∀ (a : AsBoolAlg α), f.asBoolAlg.toSupHom a = (toBoolAlg f ofBoolAlg) a
                @[simp]
                theorem RingHom.asBoolAlg_id {α : Type u_1} [BooleanRing α] :
                @[simp]
                theorem RingHom.asBoolAlg_comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} [BooleanRing α] [BooleanRing β] [BooleanRing γ] (g : β →+* γ) (f : α →+* β) :
                (g.comp f).asBoolAlg = g.asBoolAlg.comp f.asBoolAlg

                Turning a Boolean algebra into a Boolean ring #

                def AsBoolRing (α : Type u_4) :
                Type u_4

                Type synonym to view a Boolean ring as a Boolean algebra.

                Equations
                Instances For
                  def toBoolRing {α : Type u_1} :

                  The "identity" equivalence between AsBoolRing α and α.

                  Equations
                  Instances For
                    def ofBoolRing {α : Type u_1} :

                    The "identity" equivalence between α and AsBoolRing α.

                    Equations
                    Instances For
                      @[simp]
                      theorem toBoolRing_symm_eq {α : Type u_1} :
                      toBoolRing.symm = ofBoolRing
                      @[simp]
                      theorem ofBoolRing_symm_eq {α : Type u_1} :
                      ofBoolRing.symm = toBoolRing
                      @[simp]
                      theorem toBoolRing_ofBoolRing {α : Type u_1} (a : AsBoolRing α) :
                      toBoolRing (ofBoolRing a) = a
                      @[simp]
                      theorem ofBoolRing_toBoolRing {α : Type u_1} (a : α) :
                      ofBoolRing (toBoolRing a) = a
                      theorem toBoolRing_inj {α : Type u_1} {a : α} {b : α} :
                      toBoolRing a = toBoolRing b a = b
                      theorem ofBoolRing_inj {α : Type u_1} {a : AsBoolRing α} {b : AsBoolRing α} :
                      ofBoolRing a = ofBoolRing b a = b
                      Equations
                      • instInhabitedAsBoolRing = inst
                      @[reducible, inline]

                      Every generalized Boolean algebra has the structure of a non unital commutative ring with the following data:

                      • a + b unfolds to a ∆ b (symmetric difference)
                      • a * b unfolds to a ⊓ b
                      • -a unfolds to a
                      • 0 unfolds to
                      Equations
                      Instances For
                        Equations
                        • instNonUnitalCommRingAsBoolRingOfGeneralizedBooleanAlgebra = GeneralizedBooleanAlgebra.toNonUnitalCommRing
                        @[reducible, inline]

                        Every Boolean algebra has the structure of a Boolean ring with the following data:

                        • a + b unfolds to a ∆ b (symmetric difference)
                        • a * b unfolds to a ⊓ b
                        • -a unfolds to a
                        • 0 unfolds to
                        • 1 unfolds to
                        Equations
                        Instances For
                          Equations
                          • instBooleanRingAsBoolRing = BooleanAlgebra.toBooleanRing
                          @[simp]
                          theorem ofBoolRing_zero {α : Type u_1} [BooleanAlgebra α] :
                          ofBoolRing 0 =
                          @[simp]
                          theorem ofBoolRing_one {α : Type u_1} [BooleanAlgebra α] :
                          ofBoolRing 1 =
                          @[simp]
                          theorem ofBoolRing_neg {α : Type u_1} [BooleanAlgebra α] (a : AsBoolRing α) :
                          ofBoolRing (-a) = ofBoolRing a
                          @[simp]
                          theorem ofBoolRing_add {α : Type u_1} [BooleanAlgebra α] (a : AsBoolRing α) (b : AsBoolRing α) :
                          ofBoolRing (a + b) = symmDiff (ofBoolRing a) (ofBoolRing b)
                          @[simp]
                          theorem ofBoolRing_sub {α : Type u_1} [BooleanAlgebra α] (a : AsBoolRing α) (b : AsBoolRing α) :
                          ofBoolRing (a - b) = symmDiff (ofBoolRing a) (ofBoolRing b)
                          @[simp]
                          theorem ofBoolRing_mul {α : Type u_1} [BooleanAlgebra α] (a : AsBoolRing α) (b : AsBoolRing α) :
                          ofBoolRing (a * b) = ofBoolRing a ofBoolRing b
                          @[simp]
                          theorem ofBoolRing_le_ofBoolRing_iff {α : Type u_1} [BooleanAlgebra α] {a : AsBoolRing α} {b : AsBoolRing α} :
                          ofBoolRing a ofBoolRing b a * b = a
                          @[simp]
                          theorem toBoolRing_bot {α : Type u_1} [BooleanAlgebra α] :
                          toBoolRing = 0
                          @[simp]
                          theorem toBoolRing_top {α : Type u_1} [BooleanAlgebra α] :
                          toBoolRing = 1
                          @[simp]
                          theorem toBoolRing_inf {α : Type u_1} [BooleanAlgebra α] (a : α) (b : α) :
                          toBoolRing (a b) = toBoolRing a * toBoolRing b
                          @[simp]
                          theorem toBoolRing_symmDiff {α : Type u_1} [BooleanAlgebra α] (a : α) (b : α) :
                          toBoolRing (symmDiff a b) = toBoolRing a + toBoolRing b

                          Turn a bounded lattice homomorphism from Boolean algebras α to β into a ring homomorphism from α to β considered as Boolean rings.

                          Equations
                          • f.asBoolRing = { toFun := toBoolRing f ofBoolRing, map_one' := , map_mul' := , map_zero' := , map_add' := }
                          Instances For
                            @[simp]
                            theorem BoundedLatticeHom.asBoolRing_apply {α : Type u_1} {β : Type u_2} [BooleanAlgebra α] [BooleanAlgebra β] (f : BoundedLatticeHom α β) :
                            ∀ (a : AsBoolRing α), f.asBoolRing a = (toBoolRing f ofBoolRing) a
                            @[simp]
                            theorem BoundedLatticeHom.asBoolRing_comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} [BooleanAlgebra α] [BooleanAlgebra β] [BooleanAlgebra γ] (g : BoundedLatticeHom β γ) (f : BoundedLatticeHom α β) :
                            (g.comp f).asBoolRing = g.asBoolRing.comp f.asBoolRing

                            Equivalence between Boolean rings and Boolean algebras #

                            Order isomorphism between α considered as a Boolean ring considered as a Boolean algebra and α.

                            Equations
                            Instances For
                              @[simp]
                              theorem OrderIso.asBoolAlgAsBoolRing_apply (α : Type u_4) [BooleanAlgebra α] :
                              ∀ (a : AsBoolAlg (AsBoolRing α)), (OrderIso.asBoolAlgAsBoolRing α) a = ofBoolRing (ofBoolAlg a)
                              @[simp]
                              theorem OrderIso.asBoolAlgAsBoolRing_symm_apply (α : Type u_4) [BooleanAlgebra α] :
                              ∀ (a : α), (RelIso.symm (OrderIso.asBoolAlgAsBoolRing α)) a = toBoolAlg (toBoolRing a)

                              Ring isomorphism between α considered as a Boolean algebra considered as a Boolean ring and α.

                              Equations
                              Instances For
                                @[simp]
                                theorem RingEquiv.asBoolRingAsBoolAlg_apply (α : Type u_4) [BooleanRing α] :
                                ∀ (a : AsBoolRing (AsBoolAlg α)), (RingEquiv.asBoolRingAsBoolAlg α) a = ofBoolAlg (ofBoolRing a)
                                @[simp]
                                theorem RingEquiv.asBoolRingAsBoolAlg_symm_apply (α : Type u_4) [BooleanRing α] :
                                ∀ (a : α), (RingEquiv.asBoolRingAsBoolAlg α).symm a = toBoolRing (toBoolAlg a)
                                instance instOneBool :
                                Equations
                                Equations