Symmetric difference and bi-implication #
This file defines the symmetric difference and bi-implication operators in (co-)Heyting algebras.
Examples #
Some examples are
- The symmetric difference of two sets is the set of elements that are in either but not both.
- The symmetric difference on propositions is Xor'.
- The symmetric difference on BoolisBool.xor.
- The equivalence of propositions. Two propositions are equivalent if they imply each other.
- The symmetric difference translates to addition when considering a Boolean algebra as a Boolean ring.
Main declarations #
- symmDiff: The symmetric difference operator, defined as- (a \ b) ⊔ (b \ a)
- bihimp: The bi-implication operator, defined as- (b ⇨ a) ⊓ (a ⇨ b)
In generalized Boolean algebras, the symmetric difference operator is:
- symmDiff_comm: commutative, and
- symmDiff_assoc: associative.
Notations #
References #
The proof of associativity follows the note "Associativity of the Symmetric Difference of Sets: A Proof from the Book" by John McCuan:
Tags #
boolean ring, generalized boolean algebra, boolean algebra, symmetric difference, bi-implication, Heyting
Notation for symmDiff
Equations
- symmDiff.«term_∆_» = Lean.ParserDescr.trailingNode `symmDiff.«term_∆_» 100 100 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ∆ ") (Lean.ParserDescr.cat `term 101))
Instances For
Notation for bihimp
Equations
- symmDiff.«term_⇔_» = Lean.ParserDescr.trailingNode `symmDiff.«term_⇔_» 100 100 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ⇔ ") (Lean.ParserDescr.cat `term 101))
Instances For
@[simp]
@[simp]
instance
symmDiff_isCommutative
{α : Type u_2}
[GeneralizedCoheytingAlgebra α]
 :
Std.Commutative fun (x1 x2 : α) => symmDiff x1 x2
@[simp]
@[simp]
@[simp]
@[simp]
theorem
symmDiff_le
{α : Type u_2}
[GeneralizedCoheytingAlgebra α]
{a b c : α}
(ha : a ≤ b ⊔ c)
(hb : b ≤ a ⊔ c)
 :
@[simp]
theorem
Disjoint.symmDiff_eq_sup
{α : Type u_2}
[GeneralizedCoheytingAlgebra α]
{a b : α}
(h : Disjoint a b)
 :
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
instance
bihimp_isCommutative
{α : Type u_2}
[GeneralizedHeytingAlgebra α]
 :
Std.Commutative fun (x1 x2 : α) => bihimp x1 x2
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[deprecated bihimp_eq_sup_himp_inf (since := "2025-06-05")]
Alias of bihimp_eq_sup_himp_inf.
theorem
Codisjoint.bihimp_eq_inf
{α : Type u_2}
[GeneralizedHeytingAlgebra α]
{a b : α}
(h : Codisjoint a b)
 :
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
instance
symmDiff_isAssociative
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
 :
Std.Associative fun (x1 x2 : α) => symmDiff x1 x2
@[simp]
@[simp]
@[simp]
theorem
symmDiff_left_involutive
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
 :
Function.Involutive fun (x : α) => symmDiff x a
theorem
symmDiff_right_involutive
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
 :
Function.Involutive fun (x : α) => symmDiff a x
theorem
symmDiff_left_injective
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
 :
Function.Injective fun (x : α) => symmDiff x a
theorem
symmDiff_right_injective
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
 :
Function.Injective fun (x : α) => symmDiff a x
theorem
symmDiff_left_surjective
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
 :
Function.Surjective fun (x : α) => symmDiff x a
theorem
symmDiff_right_surjective
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
 :
Function.Surjective fun (x : α) => symmDiff a x
@[simp]
@[simp]
@[simp]
@[simp]
theorem
Disjoint.symmDiff_left
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
{a b c : α}
(ha : Disjoint a c)
(hb : Disjoint b c)
 :
theorem
Disjoint.symmDiff_right
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
{a b c : α}
(ha : Disjoint a b)
(hb : Disjoint a c)
 :
theorem
symmDiff_eq_iff_sdiff_eq
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
{a b c : α}
(ha : a ≤ c)
 :
CogeneralizedBooleanAlgebra isn't actually a typeclass, but the lemmas in here are dual to
the GeneralizedBooleanAlgebra ones
@[simp]
theorem
codisjoint_bihimp_sup
{α : Type u_2}
[BooleanAlgebra α]
(a b : α)
 :
Codisjoint (bihimp a b) (a ⊔ b)
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
instance
bihimp_isAssociative
{α : Type u_2}
[BooleanAlgebra α]
 :
Std.Associative fun (x1 x2 : α) => bihimp x1 x2
@[simp]
@[simp]
@[simp]
theorem
bihimp_left_involutive
{α : Type u_2}
[BooleanAlgebra α]
(a : α)
 :
Function.Involutive fun (x : α) => bihimp x a
theorem
bihimp_right_involutive
{α : Type u_2}
[BooleanAlgebra α]
(a : α)
 :
Function.Involutive fun (x : α) => bihimp a x
theorem
bihimp_left_injective
{α : Type u_2}
[BooleanAlgebra α]
(a : α)
 :
Function.Injective fun (x : α) => bihimp x a
theorem
bihimp_right_injective
{α : Type u_2}
[BooleanAlgebra α]
(a : α)
 :
Function.Injective fun (x : α) => bihimp a x
theorem
bihimp_left_surjective
{α : Type u_2}
[BooleanAlgebra α]
(a : α)
 :
Function.Surjective fun (x : α) => bihimp x a
theorem
bihimp_right_surjective
{α : Type u_2}
[BooleanAlgebra α]
(a : α)
 :
Function.Surjective fun (x : α) => bihimp a x
@[simp]
@[simp]
@[simp]
@[simp]
theorem
Codisjoint.bihimp_left
{α : Type u_2}
[BooleanAlgebra α]
{a b c : α}
(ha : Codisjoint a c)
(hb : Codisjoint b c)
 :
Codisjoint (bihimp a b) c
theorem
Codisjoint.bihimp_right
{α : Type u_2}
[BooleanAlgebra α]
{a b c : α}
(ha : Codisjoint a b)
(hb : Codisjoint a c)
 :
Codisjoint a (bihimp b c)
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
theorem
Disjoint.le_symmDiff_sup_symmDiff_left
{α : Type u_2}
[BooleanAlgebra α]
{a b c : α}
(h : Disjoint a b)
 :
theorem
Disjoint.le_symmDiff_sup_symmDiff_right
{α : Type u_2}
[BooleanAlgebra α]
{a b c : α}
(h : Disjoint b c)
 :
theorem
Codisjoint.bihimp_inf_bihimp_le_left
{α : Type u_2}
[BooleanAlgebra α]
{a b c : α}
(h : Codisjoint a b)
 :
theorem
Codisjoint.bihimp_inf_bihimp_le_right
{α : Type u_2}
[BooleanAlgebra α]
{a b c : α}
(h : Codisjoint b c)
 :
Prod #
@[simp]
theorem
symmDiff_fst
{α : Type u_2}
{β : Type u_3}
[GeneralizedCoheytingAlgebra α]
[GeneralizedCoheytingAlgebra β]
(a b : α × β)
 :
@[simp]
theorem
symmDiff_snd
{α : Type u_2}
{β : Type u_3}
[GeneralizedCoheytingAlgebra α]
[GeneralizedCoheytingAlgebra β]
(a b : α × β)
 :
@[simp]
theorem
bihimp_fst
{α : Type u_2}
{β : Type u_3}
[GeneralizedHeytingAlgebra α]
[GeneralizedHeytingAlgebra β]
(a b : α × β)
 :
@[simp]
theorem
bihimp_snd
{α : Type u_2}
{β : Type u_3}
[GeneralizedHeytingAlgebra α]
[GeneralizedHeytingAlgebra β]
(a b : α × β)
 :
Pi #
theorem
Pi.symmDiff_def
{ι : Type u_1}
{π : ι → Type u_4}
[(i : ι) → GeneralizedCoheytingAlgebra (π i)]
(a b : (i : ι) → π i)
 :
theorem
Pi.bihimp_def
{ι : Type u_1}
{π : ι → Type u_4}
[(i : ι) → GeneralizedHeytingAlgebra (π i)]
(a b : (i : ι) → π i)
 :
@[simp]
theorem
Pi.symmDiff_apply
{ι : Type u_1}
{π : ι → Type u_4}
[(i : ι) → GeneralizedCoheytingAlgebra (π i)]
(a b : (i : ι) → π i)
(i : ι)
 :
@[simp]
theorem
Pi.bihimp_apply
{ι : Type u_1}
{π : ι → Type u_4}
[(i : ι) → GeneralizedHeytingAlgebra (π i)]
(a b : (i : ι) → π i)
(i : ι)
 :