{-# OPTIONS --without-K --safe #-} open import Data.Hypergraph.Label using (HypergraphLabel) module Data.Hypergraph.Edge.Order (HL : HypergraphLabel) where import Data.List.Sort as ListSort import Data.Fin as Fin import Data.Fin.Properties as FinProp import Data.Vec as Vec import Data.Vec.Relation.Binary.Equality.Cast as VecCast import Data.Vec.Relation.Binary.Lex.Strict as Lex import Relation.Binary.PropositionalEquality as ≡ import Relation.Binary.Properties.DecTotalOrder as DTOP import Relation.Binary.Properties.StrictTotalOrder as STOP open import Data.Hypergraph.Edge HL using (Edge; ≈-Edge; ≈-Edge-IsEquivalence) open import Data.Fin using (Fin) open import Data.Nat using (ℕ; _<_) open import Data.Nat.Properties using (<-irrefl; <-trans; <-resp₂-≡; <-cmp) open import Data.Product.Base using (_,_; proj₁; proj₂) open import Data.Vec.Relation.Binary.Pointwise.Inductive using (Pointwise-≡⇒≡) open import Level using (0ℓ) open import Relation.Binary using ( Rel ; Tri ; Trichotomous ; IsStrictTotalOrder ; IsEquivalence ; _Respects_ ; DecTotalOrder ; StrictTotalOrder ) open import Relation.Nullary using (¬_) module HL = HypergraphLabel HL open HL using (Label; cast; cast-is-id) open Vec using (Vec) open ≡ using (_≡_) open HL using (_[_<_]) _<<_ : {v a : ℕ} → Rel (Vec (Fin v) a) 0ℓ _<<_ {v} = Lex.Lex-< _≡_ (Fin._<_ {v}) data <-Edge {v : ℕ} : Edge v → Edge v → Set where <-arity : {x y : Edge v} → Edge.arity x < Edge.arity y → <-Edge x y <-label : {x y : Edge v} (≡a : Edge.arity x ≡ Edge.arity y) → Edge.arity y [ cast ≡a (Edge.label x) < Edge.label y ] → <-Edge x y <-ports : {x y : Edge v} (≡a : Edge.arity x ≡ Edge.arity y) (≡l : Edge.label x HL.≈[ ≡a ] Edge.label y) → Vec.cast ≡a (Edge.ports x) << Edge.ports y → <-Edge x y <-Edge-irrefl : {v : ℕ} {x y : Edge v} → ≈-Edge x y → ¬ <-Edge x y <-Edge-irrefl record { ≡arity = ≡a } (<-arity n x≮y″ x≢y y ¬x x≮y′ x≢y y ¬x x≮y x≢y y ¬x