Documentation

Mathlib.Order.BooleanAlgebra.Defs

(Generalized) Boolean algebras #

This file sets up the theory of (generalized) Boolean algebras.

A Boolean algebra is a bounded distributive lattice with a complement operator. Boolean algebras generalize the (classical) logic of propositions and the lattice of subsets of a set.

Generalized Boolean algebras may be less familiar, but they are essentially Boolean algebras which do not necessarily have a top element () (and hence not all elements may have complements). One example in mathlib is Finset α, the type of all finite subsets of an arbitrary (not-necessarily-finite) type α.

GeneralizedBooleanAlgebra α is defined to be a distributive lattice with bottom () admitting a relative complement operator, written using "set difference" notation as x \ y (sdiff x y). For convenience, the BooleanAlgebra type class is defined to extend GeneralizedBooleanAlgebra so that it is also bundled with a \ operator.

(A terminological point: x \ y is the complement of y relative to the interval [⊥, x]. We do not yet have relative complements for arbitrary intervals, as we do not even have lattice intervals.)

Main declarations #

Implementation notes #

The sup_inf_sdiff and inf_inf_sdiff axioms for the relative complement operator in GeneralizedBooleanAlgebra are taken from Wikipedia.

Stone's paper introducing generalized Boolean algebras does not define a relative complement operator a \ b for all a, b. Instead, the postulates there amount to an assumption that for all a, b : α where a ≤ b, the equations x ⊔ a = b and x ⊓ a = ⊥ have a solution x. Disjoint.sdiff_unique proves that this x is in fact b \ a.

References #

Tags #

generalized Boolean algebras, Boolean algebras, lattices, sdiff, compl

Generalized Boolean algebras #

class GeneralizedBooleanAlgebra (α : Type u) extends DistribLattice α, SDiff α, Bot α :

A generalized Boolean algebra is a distributive lattice with and a relative complement operation \ (called sdiff, after "set difference") satisfying (a ⊓ b) ⊔ (a \ b) = a and (a ⊓ b) ⊓ (a \ b) = ⊥, i.e. a \ b is the complement of b in a.

This is a generalization of Boolean algebras which applies to Finset α for arbitrary (not-necessarily-Fintype) α.

Instances

    Boolean algebras #

    class BooleanAlgebra (α : Type u) extends DistribLattice α, HasCompl α, SDiff α, HImp α, Top α, Bot α :

    A Boolean algebra is a bounded distributive lattice with a complement operator such that x ⊓ xᶜ = ⊥ and x ⊔ xᶜ = ⊤. For convenience, it must also provide a set difference operation \ and a Heyting implication satisfying x \ y = x ⊓ yᶜ and x ⇨ y = y ⊔ xᶜ.

    This is a generalization of (classical) logic of propositions, or the powerset lattice.

    Since BoundedOrder, OrderBot, and OrderTop are mixins that require LE to be present at define-time, the extends mechanism does not work with them. Instead, we extend using the underlying Bot and Top data typeclasses, and replicate the order axioms of those classes here. A "forgetful" instance back to BoundedOrder is provided.

    Instances
      @[instance 100]
      Equations
      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.
      theorem Bool.sup_eq_bor :
      (fun (x1 x2 : Bool) => max x1 x2) = or
      theorem Bool.inf_eq_band :
      (fun (x1 x2 : Bool) => min x1 x2) = and
      Equations
      • One or more equations did not get rendered due to their size.

      An alternative constructor for boolean algebras: a distributive lattice that is complemented is a boolean algebra.

      This is not an instance, because it creates data using choice.

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