Basic properties of Boolean algebras #
This file provides some basic definitions, functions as well as lemmas for functions and type
classes related to Boolean algebras as defined in Mathlib/Order/BooleanAlgebra/Defs.lean
.
References #
- https://en.wikipedia.org/wiki/Boolean_algebra_(structure)#Generalizations
- [Postulates for Boolean Algebras and Generalized Boolean Algebras, M.H. Stone][Stone1935]
- [Lattice Theory: Foundation, George Grätzer][Gratzer2011]
Tags #
generalized Boolean algebras, Boolean algebras, lattices, sdiff, compl
Generalized Boolean algebras #
Some of the lemmas in this section are from:
- [Lattice Theory: Foundation, George Grätzer][Gratzer2011]
- https://ncatlab.org/nlab/show/relative+complement
- https://people.math.gatech.edu/~mccuan/courses/4317/symmetricdifference.pdf
@[simp]
@[simp]
@[simp]
@[simp]
@[instance 100]
instance
GeneralizedBooleanAlgebra.toOrderBot
{α : Type u}
[GeneralizedBooleanAlgebra α]
:
OrderBot α
Equations
- GeneralizedBooleanAlgebra.toOrderBot = { toBot := inst✝.toBot, bot_le := ⋯ }
theorem
sdiff_unique
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(s : x ⊓ y ⊔ z = x)
(i : x ⊓ y ⊓ z = ⊥)
:
@[simp]
@[simp]
@[simp]
@[instance 100]
instance
GeneralizedBooleanAlgebra.toGeneralizedCoheytingAlgebra
{α : Type u}
[GeneralizedBooleanAlgebra α]
:
@[simp]
theorem
Disjoint.sdiff_eq_of_sup_eq
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(hi : Disjoint x z)
(hs : x ⊔ z = y)
:
theorem
Disjoint.sdiff_unique
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(hd : Disjoint x z)
(hz : z ≤ y)
(hs : y ≤ x ⊔ z)
:
theorem
disjoint_sdiff_iff_le
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(hz : z ≤ y)
(hx : x ≤ y)
:
theorem
le_iff_disjoint_sdiff
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(hz : z ≤ y)
(hx : x ≤ y)
:
theorem
le_iff_eq_sup_sdiff
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(hz : z ≤ y)
(hx : x ≤ y)
:
@[simp]
theorem
sdiff_lt_sdiff_right
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(h : x < y)
(hz : z ≤ x)
:
theorem
sdiff_sdiff_eq_sdiff_sup
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(h : z ≤ x)
:
@[simp]
theorem
sdiff_eq_symm
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(hy : y ≤ x)
(h : x \ y = z)
:
theorem
eq_of_sdiff_eq_sdiff
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(hxz : x ≤ z)
(hyz : y ≤ z)
(h : z \ x = z \ y)
:
theorem
sdiff_sdiff_sdiff_cancel_left
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(hca : z ≤ x)
:
theorem
sdiff_sdiff_sdiff_cancel_right
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(hcb : z ≤ y)
:
See also sdiff_inf_right_comm
.
See also inf_sdiff_assoc
.
@[deprecated sdiff_inf_right_comm (since := "2025-01-08")]
Alias of sdiff_inf_right_comm
.
See also inf_sdiff_assoc
.
theorem
sup_lt_of_lt_sdiff_left
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(h : y < z \ x)
(hxz : x ≤ z)
:
theorem
sup_lt_of_lt_sdiff_right
{α : Type u}
{x y z : α}
[GeneralizedBooleanAlgebra α]
(h : x < z \ y)
(hyz : y ≤ z)
:
instance
Prod.instGeneralizedBooleanAlgebra
{α : Type u}
{β : Type u_1}
[GeneralizedBooleanAlgebra α]
[GeneralizedBooleanAlgebra β]
:
GeneralizedBooleanAlgebra (α × β)
Equations
- Prod.instGeneralizedBooleanAlgebra = { toDistribLattice := Prod.instDistribLattice α β, toSDiff := Prod.instSDiff α β, toBot := Prod.instBot α β, sup_inf_sdiff := ⋯, inf_inf_sdiff := ⋯ }
instance
Pi.instGeneralizedBooleanAlgebra
{ι : Type u_2}
{α : ι → Type u_3}
[(i : ι) → GeneralizedBooleanAlgebra (α i)]
:
GeneralizedBooleanAlgebra ((i : ι) → α i)
Equations
- Pi.instGeneralizedBooleanAlgebra = { toDistribLattice := Pi.instDistribLattice, toSDiff := Pi.sdiff, toBot := Pi.instBotForall, sup_inf_sdiff := ⋯, inf_inf_sdiff := ⋯ }
Boolean algebras #
@[reducible, inline]
abbrev
GeneralizedBooleanAlgebra.toBooleanAlgebra
{α : Type u}
[GeneralizedBooleanAlgebra α]
[OrderTop α]
:
A bounded generalized boolean algebra is a boolean algebra.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[instance 100]
@[instance 100]
Equations
- BooleanAlgebra.toGeneralizedBooleanAlgebra = { toDistribLattice := inst✝.toDistribLattice, toSDiff := inst✝.toSDiff, toBot := inst✝.toBot, sup_inf_sdiff := ⋯, inf_inf_sdiff := ⋯ }
@[instance 100]
Equations
- One or more equations did not get rendered due to their size.
@[simp]
@[simp]
@[simp]
Equations
- One or more equations did not get rendered due to their size.
@[simp]
instance
Prod.instBooleanAlgebra
{α : Type u}
{β : Type u_1}
[BooleanAlgebra α]
[BooleanAlgebra β]
:
BooleanAlgebra (α × β)
Equations
- One or more equations did not get rendered due to their size.
instance
Pi.instBooleanAlgebra
{ι : Type u}
{α : ι → Type v}
[(i : ι) → BooleanAlgebra (α i)]
:
BooleanAlgebra ((i : ι) → α i)
Equations
- One or more equations did not get rendered due to their size.
@[reducible, inline]
abbrev
Function.Injective.generalizedBooleanAlgebra
{α : Type u}
{β : Type u_1}
[Max α]
[Min α]
[Bot α]
[SDiff α]
[GeneralizedBooleanAlgebra β]
(f : α → β)
(hf : Injective f)
(map_sup : ∀ (a b : α), f (a ⊔ b) = f a ⊔ f b)
(map_inf : ∀ (a b : α), f (a ⊓ b) = f a ⊓ f b)
(map_bot : f ⊥ = ⊥)
(map_sdiff : ∀ (a b : α), f (a \ b) = f a \ f b)
:
Pullback a GeneralizedBooleanAlgebra
along an injection.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[reducible, inline]
abbrev
Function.Injective.booleanAlgebra
{α : Type u}
{β : Type u_1}
[Max α]
[Min α]
[Top α]
[Bot α]
[HasCompl α]
[SDiff α]
[HImp α]
[BooleanAlgebra β]
(f : α → β)
(hf : Injective f)
(map_sup : ∀ (a b : α), f (a ⊔ b) = f a ⊔ f b)
(map_inf : ∀ (a b : α), f (a ⊓ b) = f a ⊓ f b)
(map_top : f ⊤ = ⊤)
(map_bot : f ⊥ = ⊥)
(map_compl : ∀ (a : α), f aᶜ = (f a)ᶜ)
(map_sdiff : ∀ (a b : α), f (a \ b) = f a \ f b)
(map_himp : ∀ (a b : α), f (a ⇨ b) = f a ⇨ f b)
:
Pullback a BooleanAlgebra
along an injection.
Equations
- One or more equations did not get rendered due to their size.