{-# OPTIONS --cubical-compatible --safe #-}
module Relation.Binary.Reasoning.MultiSetoid where
open import Level using (Level; _⊔_)
open import Function.Base using (case_of_)
open import Relation.Binary.Core using (_⇒_)
open import Relation.Binary.Definitions using (Trans; Reflexive)
open import Relation.Binary.Bundles using (Setoid)
open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
open import Relation.Binary.Reasoning.Syntax
private
variable
a ℓ : Level
module _ (S : Setoid a ℓ) where
open Setoid S
data IsRelatedTo (x y : _) : Set (a ⊔ ℓ) where
relTo : (x≈y : x ≈ y) → IsRelatedTo x y
start : IsRelatedTo ⇒ _≈_
start (relTo x≈y) = x≈y
≡-go : Trans _≡_ IsRelatedTo IsRelatedTo
≡-go x≡y (relTo y∼z) = relTo (case x≡y of λ where ≡.refl → y∼z)
≈-go : Trans _≈_ IsRelatedTo IsRelatedTo
≈-go x≈y (relTo y≈z) = relTo (trans x≈y y≈z)
end : Reflexive IsRelatedTo
end = relTo refl
open begin-syntax IsRelatedTo start public
renaming (begin_ to begin⟨_⟩_)
module _ {S : Setoid a ℓ} where
open Setoid S
open ≡-syntax (IsRelatedTo S) (≡-go S)
open ≈-syntax (IsRelatedTo S) (IsRelatedTo S) (≈-go S) sym public
open end-syntax (IsRelatedTo S) (end S) public