{-# OPTIONS --cubical-compatible --safe #-}
{-# OPTIONS --postfix-projections #-}
open import Relation.Binary.Core using (Rel)
open import Relation.Binary.Bundles using (Setoid)
module Data.List.Relation.Binary.Sublist.Setoid
{c ℓ} (S : Setoid c ℓ) where
open import Level using (_⊔_)
open import Data.List.Base using (List; []; _∷_)
import Data.List.Relation.Binary.Equality.Setoid as SetoidEquality
import Data.List.Relation.Binary.Sublist.Heterogeneous as Heterogeneous
import Data.List.Relation.Binary.Sublist.Heterogeneous.Core
as HeterogeneousCore
import Data.List.Relation.Binary.Sublist.Heterogeneous.Properties
as HeterogeneousProperties
open import Data.Product.Base using (∃; ∃₂; _×_; _,_; proj₂)
open import Relation.Binary.Core using (_⇒_)
open import Relation.Binary.Bundles using (Preorder; Poset)
open import Relation.Binary.Structures using (IsPreorder; IsPartialOrder)
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
open import Relation.Nullary using (¬_; Dec; yes; no)
open Setoid S renaming (Carrier to A)
open SetoidEquality S
infix 4 _⊆_ _⊇_ _⊂_ _⊃_ _⊈_ _⊉_ _⊄_ _⊅_
_⊆_ : Rel (List A) (c ⊔ ℓ)
_⊆_ = Heterogeneous.Sublist _≈_
_⊇_ : Rel (List A) (c ⊔ ℓ)
xs ⊇ ys = ys ⊆ xs
_⊂_ : Rel (List A) (c ⊔ ℓ)
xs ⊂ ys = xs ⊆ ys × ¬ (xs ≋ ys)
_⊃_ : Rel (List A) (c ⊔ ℓ)
xs ⊃ ys = ys ⊂ xs
_⊈_ : Rel (List A) (c ⊔ ℓ)
xs ⊈ ys = ¬ (xs ⊆ ys)
_⊉_ : Rel (List A) (c ⊔ ℓ)
xs ⊉ ys = ¬ (xs ⊇ ys)
_⊄_ : Rel (List A) (c ⊔ ℓ)
xs ⊄ ys = ¬ (xs ⊂ ys)
_⊅_ : Rel (List A) (c ⊔ ℓ)
xs ⊅ ys = ¬ (xs ⊃ ys)
open HeterogeneousCore _≈_ public
using ([]; _∷_; _∷ʳ_)
open Heterogeneous {R = _≈_} public
hiding (Sublist; []; _∷_; _∷ʳ_)
renaming
( toAny to to∈
; fromAny to from∈
)
open Disjoint public
using ([])
open DisjointUnion public
using ([])
⊆-reflexive : _≋_ ⇒ _⊆_
⊆-reflexive = HeterogeneousProperties.fromPointwise
open HeterogeneousProperties.Reflexivity {R = _≈_} refl public using ()
renaming (refl to ⊆-refl)
open HeterogeneousProperties.Transitivity {R = _≈_} {S = _≈_} {T = _≈_} trans public using ()
renaming (trans to ⊆-trans)
open HeterogeneousProperties.Antisymmetry {R = _≈_} {S = _≈_} (λ x≈y _ → x≈y) public using ()
renaming (antisym to ⊆-antisym)
⊆-isPreorder : IsPreorder _≋_ _⊆_
⊆-isPreorder = record
{ isEquivalence = ≋-isEquivalence
; reflexive = ⊆-reflexive
; trans = ⊆-trans
}
⊆-isPartialOrder : IsPartialOrder _≋_ _⊆_
⊆-isPartialOrder = record
{ isPreorder = ⊆-isPreorder
; antisym = ⊆-antisym
}
⊆-preorder : Preorder c (c ⊔ ℓ) (c ⊔ ℓ)
⊆-preorder = record
{ isPreorder = ⊆-isPreorder
}
⊆-poset : Poset c (c ⊔ ℓ) (c ⊔ ℓ)
⊆-poset = record
{ isPartialOrder = ⊆-isPartialOrder
}
record RawPushout {xs ys zs : List A} (τ : xs ⊆ ys) (σ : xs ⊆ zs) : Set (c ⊔ ℓ) where
field
{upperBound} : List A
leg₁ : ys ⊆ upperBound
leg₂ : zs ⊆ upperBound
open RawPushout
infixr 5 _∷ʳ₁_ _∷ʳ₂_
_∷ʳ₁_ : ∀ {xs ys zs : List A} {τ : xs ⊆ ys} {σ : xs ⊆ zs} →
(y : A) → RawPushout τ σ → RawPushout (y ∷ʳ τ) σ
y ∷ʳ₁ rpo = record
{ leg₁ = refl ∷ leg₁ rpo
; leg₂ = y ∷ʳ leg₂ rpo
}
_∷ʳ₂_ : ∀ {xs ys zs : List A} {τ : xs ⊆ ys} {σ : xs ⊆ zs} →
(z : A) → RawPushout τ σ → RawPushout τ (z ∷ʳ σ)
z ∷ʳ₂ rpo = record
{ leg₁ = z ∷ʳ leg₁ rpo
; leg₂ = refl ∷ leg₂ rpo
}
∷-rpo : ∀ {x y z : A} {xs ys zs : List A} {τ : xs ⊆ ys} {σ : xs ⊆ zs} →
(x≈y : x ≈ y) (x≈z : x ≈ z) → RawPushout τ σ → RawPushout (x≈y ∷ τ) (x≈z ∷ σ)
∷-rpo x≈y x≈z rpo = record
{ leg₁ = sym x≈y ∷ leg₁ rpo
; leg₂ = sym x≈z ∷ leg₂ rpo
}
⊆-pushoutˡ : ∀ {xs ys zs : List A} →
(τ : xs ⊆ ys) (σ : xs ⊆ zs) → RawPushout τ σ
⊆-pushoutˡ [] σ = record { leg₁ = σ ; leg₂ = ⊆-refl }
⊆-pushoutˡ (y ∷ʳ τ) σ = y ∷ʳ₁ ⊆-pushoutˡ τ σ
⊆-pushoutˡ τ@(_ ∷ _) (z ∷ʳ σ) = z ∷ʳ₂ ⊆-pushoutˡ τ σ
⊆-pushoutˡ (x≈y ∷ τ) (x≈z ∷ σ) = ∷-rpo x≈y x≈z (⊆-pushoutˡ τ σ)
⊆-joinˡ : ∀ {xs ys zs : List A} →
(τ : xs ⊆ ys) (σ : xs ⊆ zs) → ∃ λ us → xs ⊆ us
⊆-joinˡ τ σ = upperBound rpo , ⊆-trans τ (leg₁ rpo)
where
rpo = ⊆-pushoutˡ τ σ
record UpperBound {xs ys zs} (τ : xs ⊆ zs) (σ : ys ⊆ zs) : Set (c ⊔ ℓ) where
field
{theUpperBound} : List A
sub : theUpperBound ⊆ zs
inj₁ : xs ⊆ theUpperBound
inj₂ : ys ⊆ theUpperBound
open UpperBound
infixr 5 _∷ₗ-ub_ _∷ᵣ-ub_
∷ₙ-ub : ∀ {xs ys zs} {τ : xs ⊆ zs} {σ : ys ⊆ zs} {x} →
UpperBound τ σ → UpperBound (x ∷ʳ τ) (x ∷ʳ σ)
∷ₙ-ub u = record
{ sub = _ ∷ʳ u .sub
; inj₁ = u .inj₁
; inj₂ = u .inj₂
}
_∷ₗ-ub_ : ∀ {xs ys zs} {τ : xs ⊆ zs} {σ : ys ⊆ zs} {x y} →
(x≈y : x ≈ y) → UpperBound τ σ → UpperBound (x≈y ∷ τ) (y ∷ʳ σ)
x≈y ∷ₗ-ub u = record
{ sub = refl ∷ u .sub
; inj₁ = x≈y ∷ u .inj₁
; inj₂ = _ ∷ʳ u .inj₂
}
_∷ᵣ-ub_ : ∀ {xs ys zs} {τ : xs ⊆ zs} {σ : ys ⊆ zs} {x y} →
(x≈y : x ≈ y) → UpperBound τ σ → UpperBound (y ∷ʳ τ) (x≈y ∷ σ)
x≈y ∷ᵣ-ub u = record
{ sub = refl ∷ u .sub
; inj₁ = _ ∷ʳ u .inj₁
; inj₂ = x≈y ∷ u .inj₂
}
⊆-disjoint-union : ∀ {xs ys zs} {τ : xs ⊆ zs} {σ : ys ⊆ zs} →
Disjoint τ σ → UpperBound τ σ
⊆-disjoint-union [] = record { sub = [] ; inj₁ = [] ; inj₂ = [] }
⊆-disjoint-union (x ∷ₙ d) = ∷ₙ-ub (⊆-disjoint-union d)
⊆-disjoint-union (x≈y ∷ₗ d) = x≈y ∷ₗ-ub (⊆-disjoint-union d)
⊆-disjoint-union (x≈y ∷ᵣ d) = x≈y ∷ᵣ-ub (⊆-disjoint-union d)