------------------------------------------------------------------------ -- The Agda standard library -- -- Pointwise products of binary relations ------------------------------------------------------------------------ module Relation.Binary.Product.Pointwise where open import Data.Product as Prod open import Data.Sum open import Data.Unit using (⊤) open import Function open import Function.Equality as F using (_⟶_; _⟨$⟩_) open import Function.Equivalence as Eq using (Equivalence; _⇔_; module Equivalence) open import Function.Injection as Inj using (Injection; _↣_; module Injection) open import Function.Inverse as Inv using (Inverse; _↔_; module Inverse) open import Function.LeftInverse as LeftInv using (LeftInverse; _↞_; _LeftInverseOf_; module LeftInverse) open import Function.Related open import Function.Surjection as Surj using (Surjection; _↠_; module Surjection) open import Level import Relation.Nullary.Decidable as Dec open import Relation.Nullary.Product open import Relation.Binary open import Relation.Binary.PropositionalEquality as P using (_≡_) module _ {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where infixr 2 _×-Rel_ _×-Rel_ : Rel A₁ ℓ₁ → Rel A₂ ℓ₂ → Rel (A₁ × A₂) _ _∼₁_ ×-Rel _∼₂_ = (_∼₁_ on proj₁) -×- (_∼₂_ on proj₂) -- Some properties which are preserved by ×-Rel (under certain -- assumptions). _×-reflexive_ : ∀ {_≈₁_ _∼₁_ _≈₂_ _∼₂_} → _≈₁_ ⇒ _∼₁_ → _≈₂_ ⇒ _∼₂_ → (_≈₁_ ×-Rel _≈₂_) ⇒ (_∼₁_ ×-Rel _∼₂_) refl₁ ×-reflexive refl₂ = λ x≈y → (refl₁ (proj₁ x≈y) , refl₂ (proj₂ x≈y)) _×-refl_ : ∀ {_∼₁_ _∼₂_} → Reflexive _∼₁_ → Reflexive _∼₂_ → Reflexive (_∼₁_ ×-Rel _∼₂_) refl₁ ×-refl refl₂ = (refl₁ , refl₂) ×-irreflexive₁ : ∀ {_≈₁_ _<₁_ _≈₂_ _<₂_} → Irreflexive _≈₁_ _<₁_ → Irreflexive (_≈₁_ ×-Rel _≈₂_) (_<₁_ ×-Rel _<₂_) ×-irreflexive₁ ir = λ x≈y x