From ee5c56a58154840f08a18d581983d6f6f3630970 Mon Sep 17 00:00:00 2001 From: Jacques Comeaux Date: Wed, 9 Jul 2025 12:15:51 -0500 Subject: Split hypergraph label and edge into separate files --- Data/Hypergraph/Edge.agda | 326 ++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 326 insertions(+) create mode 100644 Data/Hypergraph/Edge.agda (limited to 'Data/Hypergraph/Edge.agda') diff --git a/Data/Hypergraph/Edge.agda b/Data/Hypergraph/Edge.agda new file mode 100644 index 0000000..3c2a3a1 --- /dev/null +++ b/Data/Hypergraph/Edge.agda @@ -0,0 +1,326 @@ +{-# OPTIONS --without-K --safe #-} + +open import Data.Hypergraph.Label using (HypergraphLabel) + +module Data.Hypergraph.Edge (HL : HypergraphLabel) where + + +open import Relation.Binary using (Rel; IsStrictTotalOrder; Tri; Trichotomous; _Respects_) +open import Data.Castable using (IsCastable) +open import Data.Fin using (Fin) +open import Data.Fin.Show using () renaming (show to showFin) +open import Data.Nat.Base using (ℕ; _<_) +open import Data.Nat.Properties using (<-irrefl; <-trans; <-resp₂-≡; <-cmp) +open import Data.Product.Base using (_,_; proj₁; proj₂) +open import Data.String using (String; _<+>_) +open import Data.Vec.Relation.Binary.Pointwise.Inductive using (≡⇒Pointwise-≡; Pointwise-≡⇒≡) +open import Data.Vec.Show using () renaming (show to showVec) +open import Level using (0ℓ) +open import Relation.Binary.Bundles using (DecTotalOrder; StrictTotalOrder) +open import Relation.Binary.Structures using (IsEquivalence) +open import Relation.Nullary using (¬_) + +import Data.Fin.Base as Fin +import Data.Fin.Properties as FinProp +import Data.Vec.Base as VecBase +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 + +module HL = HypergraphLabel HL +open HL using (Label; cast; cast-is-id) +open VecBase using (Vec) + +record Edge (v : ℕ) : Set where + field + {arity} : ℕ + label : Label arity + ports : Vec (Fin v) arity + +open ≡ using (_≡_) +open VecCast using (_≈[_]_) + +record ≈-Edge {n : ℕ} (E E′ : Edge n) : Set where + module E = Edge E + module E′ = Edge E′ + field + ≡arity : E.arity ≡ E′.arity + ≡label : cast ≡arity E.label ≡ E′.label + ≡ports : E.ports ≈[ ≡arity ] E′.ports + +≈-Edge-refl : {v : ℕ} {x : Edge v} → ≈-Edge x x +≈-Edge-refl {_} {x} = record + { ≡arity = ≡.refl + ; ≡label = HL.≈-reflexive ≡.refl + ; ≡ports = VecCast.≈-reflexive ≡.refl + } + where + open Edge x using (arity; label) + open DecTotalOrder (HL.decTotalOrder arity) using (module Eq) + +≈-Edge-sym : {v : ℕ} {x y : Edge v} → ≈-Edge x y → ≈-Edge y x +≈-Edge-sym {_} {x} {y} x≈y = record + { ≡arity = ≡.sym ≡arity + ; ≡label = HL.≈-sym ≡label + ; ≡ports = VecCast.≈-sym ≡ports + } + where + open ≈-Edge x≈y + open DecTotalOrder (HL.decTotalOrder E.arity) using (module Eq) + +≈-Edge-trans : {v : ℕ} {i j k : Edge v} → ≈-Edge i j → ≈-Edge j k → ≈-Edge i k +≈-Edge-trans {_} {i} {j} {k} i≈j j≈k = record + { ≡arity = ≡.trans i≈j.≡arity j≈k.≡arity + ; ≡label = HL.≈-trans i≈j.≡label j≈k.≡label + ; ≡ports = VecCast.≈-trans i≈j.≡ports j≈k.≡ports + } + where + module i≈j = ≈-Edge i≈j + module j≈k = ≈-Edge j≈k + +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) + → VecBase.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 showVec showFin p + +open module STOP′ {v} = STOP (strictTotalOrder {v}) using (decTotalOrder) public -- cgit v1.2.3