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
Bool
isBool.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, andsymmDiff_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
Equations
- ⋯ = ⋯
@[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
Equations
- ⋯ = ⋯
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
theorem
Codisjoint.bihimp_eq_inf
{α : Type u_2}
[GeneralizedHeytingAlgebra α]
{a : α}
{b : α}
(h : Codisjoint a b)
:
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
theorem
IsCompl.symmDiff_eq_top
{α : Type u_2}
[CoheytingAlgebra α]
{a : α}
{b : α}
(h : IsCompl a b)
:
@[simp]
theorem
inf_symmDiff_distrib_left
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
(b : α)
(c : α)
:
theorem
inf_symmDiff_distrib_right
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
(b : α)
(c : α)
:
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
instance
symmDiff_isAssociative
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
:
Std.Associative fun (x1 x2 : α) => symmDiff x1 x2
Equations
- ⋯ = ⋯
@[simp]
theorem
symmDiff_symmDiff_cancel_left
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
(b : α)
:
@[simp]
theorem
symmDiff_symmDiff_cancel_right
{α : Type u_2}
[GeneralizedBooleanAlgebra α]
(a : α)
(b : α)
:
@[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]
theorem
bihimp_eq_inf
{α : Type u_2}
[BooleanAlgebra α]
(a : α)
(b : α)
:
bihimp a b = a ⊓ b ↔ Codisjoint a b
@[simp]
theorem
bihimp_le_iff_left
{α : Type u_2}
[BooleanAlgebra α]
(a : α)
(b : α)
:
bihimp a b ≤ a ↔ Codisjoint a b
@[simp]
theorem
bihimp_le_iff_right
{α : Type u_2}
[BooleanAlgebra α]
(a : α)
(b : α)
:
bihimp a b ≤ b ↔ Codisjoint a b
instance
bihimp_isAssociative
{α : Type u_2}
[BooleanAlgebra α]
:
Std.Associative fun (x1 x2 : α) => bihimp x1 x2
Equations
- ⋯ = ⋯
@[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 : (i : ι) → π i)
(b : (i : ι) → π i)
:
theorem
Pi.bihimp_def
{ι : Type u_1}
{π : ι → Type u_4}
[(i : ι) → GeneralizedHeytingAlgebra (π i)]
(a : (i : ι) → π i)
(b : (i : ι) → π i)
:
@[simp]
theorem
Pi.symmDiff_apply
{ι : Type u_1}
{π : ι → Type u_4}
[(i : ι) → GeneralizedCoheytingAlgebra (π i)]
(a : (i : ι) → π i)
(b : (i : ι) → π i)
(i : ι)
:
@[simp]
theorem
Pi.bihimp_apply
{ι : Type u_1}
{π : ι → Type u_4}
[(i : ι) → GeneralizedHeytingAlgebra (π i)]
(a : (i : ι) → π i)
(b : (i : ι) → π i)
(i : ι)
: