@@ -627,20 +627,7 @@ true <? _ = no (λ())
627627 }
628628
629629------------------------------------------------------------------------
630- -- Properties of _xor_
631-
632- xor-is-ok : ∀ x y → x xor y ≡ (x ∨ y) ∧ not (x ∧ y)
633- xor-is-ok true y = refl
634- xor-is-ok false y = sym (∧-identityʳ _)
635-
636- xor-∧-commutativeRing : CommutativeRing 0ℓ 0ℓ
637- xor-∧-commutativeRing = ⊕-∧-commutativeRing
638- where
639- open BooleanAlgebraProperties ∨-∧-booleanAlgebra
640- open XorRing _xor_ xor-is-ok
641-
642- ------------------------------------------------------------------------
643- -- Miscellaneous other properties
630+ -- Properties of not
644631
645632not-involutive : Involutive not
646633not-involutive true = refl
@@ -660,6 +647,84 @@ not-¬ {false} refl ()
660647¬-not {false} {true} _ = refl
661648¬-not {false} {false} x≢y = ⊥-elim (x≢y refl)
662649
650+ ------------------------------------------------------------------------
651+ -- Properties of _xor_
652+
653+ xor-is-ok : ∀ x y → x xor y ≡ (x ∨ y) ∧ not (x ∧ y)
654+ xor-is-ok true y = refl
655+ xor-is-ok false y = sym (∧-identityʳ _)
656+
657+ true-xor : ∀ x → true xor x ≡ not x
658+ true-xor false = refl
659+ true-xor true = refl
660+
661+ xor-same : ∀ x → x xor x ≡ false
662+ xor-same false = refl
663+ xor-same true = refl
664+
665+ not-distribˡ-xor : ∀ x y → not (x xor y) ≡ (not x) xor y
666+ not-distribˡ-xor false y = refl
667+ not-distribˡ-xor true y = not-involutive _
668+
669+ not-distribʳ-xor : ∀ x y → not (x xor y) ≡ x xor (not y)
670+ not-distribʳ-xor false y = refl
671+ not-distribʳ-xor true y = refl
672+
673+ xor-assoc : Associative _xor_
674+ xor-assoc true y z = sym (not-distribˡ-xor y z)
675+ xor-assoc false y z = refl
676+
677+ xor-comm : Commutative _xor_
678+ xor-comm false false = refl
679+ xor-comm false true = refl
680+ xor-comm true false = refl
681+ xor-comm true true = refl
682+
683+ xor-identityˡ : LeftIdentity false _xor_
684+ xor-identityˡ _ = refl
685+
686+ xor-identityʳ : RightIdentity false _xor_
687+ xor-identityʳ false = refl
688+ xor-identityʳ true = refl
689+
690+ xor-identity : Identity false _xor_
691+ xor-identity = xor-identityˡ , xor-identityʳ
692+
693+ xor-inverseˡ : LeftInverse true not _xor_
694+ xor-inverseˡ false = refl
695+ xor-inverseˡ true = refl
696+
697+ xor-inverseʳ : RightInverse true not _xor_
698+ xor-inverseʳ x = xor-comm x (not x) ⟨ trans ⟩ xor-inverseˡ x
699+
700+ xor-inverse : Inverse true not _xor_
701+ xor-inverse = xor-inverseˡ , xor-inverseʳ
702+
703+ ∧-distribˡ-xor : _∧_ DistributesOverˡ _xor_
704+ ∧-distribˡ-xor false y z = refl
705+ ∧-distribˡ-xor true y z = refl
706+
707+ ∧-distribʳ-xor : _∧_ DistributesOverʳ _xor_
708+ ∧-distribʳ-xor x false z = refl
709+ ∧-distribʳ-xor x true false = sym (xor-identityʳ x)
710+ ∧-distribʳ-xor x true true = sym (xor-same x)
711+
712+ ∧-distrib-xor : _∧_ DistributesOver _xor_
713+ ∧-distrib-xor = ∧-distribˡ-xor , ∧-distribʳ-xor
714+
715+ xor-annihilates-not : ∀ x y → (not x) xor (not y) ≡ x xor y
716+ xor-annihilates-not false y = not-involutive _
717+ xor-annihilates-not true y = refl
718+
719+ xor-∧-commutativeRing : CommutativeRing 0ℓ 0ℓ
720+ xor-∧-commutativeRing = ⊕-∧-commutativeRing
721+ where
722+ open BooleanAlgebraProperties ∨-∧-booleanAlgebra
723+ open XorRing _xor_ xor-is-ok
724+
725+ ------------------------------------------------------------------------
726+ -- Miscellaneous other properties
727+
663728⇔→≡ : {x y z : Bool} → x ≡ z ⇔ y ≡ z → x ≡ y
664729⇔→≡ {true } {true } hyp = refl
665730⇔→≡ {true } {false} {true } hyp = sym (Equivalence.to hyp ⟨$⟩ refl)
0 commit comments