From 298baf2b69620106e95b52206e02d58ad8cb9fc8 Mon Sep 17 00:00:00 2001 From: Jacques Comeaux Date: Mon, 3 Nov 2025 23:29:21 -0600 Subject: Use permutation for equivalence of hypergraphs --- Data/Hypergraph/Edge.agda | 368 ++++++++-------------------------------------- 1 file changed, 58 insertions(+), 310 deletions(-) (limited to 'Data/Hypergraph/Edge.agda') diff --git a/Data/Hypergraph/Edge.agda b/Data/Hypergraph/Edge.agda index ee32393..1e24559 100644 --- a/Data/Hypergraph/Edge.agda +++ b/Data/Hypergraph/Edge.agda @@ -4,43 +4,32 @@ open import Data.Hypergraph.Label using (HypergraphLabel) module Data.Hypergraph.Edge (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 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 using (ℕ; _<_) -open import Data.Nat.Properties using (<-irrefl; <-trans; <-resp₂-≡; <-cmp) -open import Data.Product.Base using (_,_; proj₁; proj₂) +open import Data.Nat using (ℕ) 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 using (Setoid; DecTotalOrder; StrictTotalOrder; IsEquivalence) -open import Relation.Nullary using (¬_) - +open import Relation.Binary using (Setoid; IsEquivalence) module HL = HypergraphLabel HL + open HL using (Label; cast; cast-is-id) open Vec using (Vec) record Edge (v : ℕ) : Set where + constructor mkEdge field {arity} : ℕ label : Label arity ports : Vec (Fin v) arity map : {n m : ℕ} → (Fin n → Fin m) → Edge n → Edge m -map {n} {m} f edge = record +map f edge = record { label = label ; ports = Vec.map f ports } @@ -50,299 +39,58 @@ map {n} {m} f edge = record 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 - -≈-Edge-IsEquivalence : {v : ℕ} → IsEquivalence (≈-Edge {v}) -≈-Edge-IsEquivalence = record - { refl = ≈-Edge-refl - ; sym = ≈-Edge-sym - ; trans = ≈-Edge-trans - } - -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 showVec showFin p - -open module STOP′ {v} = STOP (strictTotalOrder {v}) using (decTotalOrder) public - -≈-Edge⇒≡ : {v : ℕ} {x y : Edge v} → ≈-Edge x y → x ≡ y -≈-Edge⇒≡ {v} {record { label = l ; ports = p }} record { ≡arity = ≡.refl ; ≡label = ≡.refl ; ≡ports = ≡.refl } - rewrite cast-is-id ≡.refl l - rewrite VecCast.cast-is-id ≡.refl p = ≡.refl - -module Sort {v} = ListSort (decTotalOrder {v}) -open Sort using (sort) public +module _ {v : ℕ} where + + -- an equivalence relation on edges with v nodes + record _≈_ (E E′ : Edge v) : Set where + constructor mk≈ + 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 + + ≈-refl : {x : Edge v} → x ≈ x + ≈-refl = record + { ≡arity = ≡.refl + ; ≡label = HL.≈-reflexive ≡.refl + ; ≡ports = VecCast.≈-reflexive ≡.refl + } + + ≈-sym : {x y : Edge v} → x ≈ y → y ≈ x + ≈-sym x≈y = record + { ≡arity = ≡.sym ≡arity + ; ≡label = HL.≈-sym ≡label + ; ≡ports = VecCast.≈-sym ≡ports + } + where + open _≈_ x≈y + + ≈-trans : {i j k : Edge v} → i ≈ j → j ≈ k → i ≈ k + ≈-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 = _≈_ i≈j + module j≈k = _≈_ j≈k + + ≈-IsEquivalence : IsEquivalence _≈_ + ≈-IsEquivalence = record + { refl = ≈-refl + ; sym = ≈-sym + ; trans = ≈-trans + } + + show : Edge v → String + show (mkEdge {a} l p) = HL.showLabel a l <+> showVec showFin p + + ≈⇒≡ : {x y : Edge v} → x ≈ y → x ≡ y + ≈⇒≡ {mkEdge l p} (mk≈ ≡.refl ≡.refl ≡.refl) + rewrite cast-is-id ≡.refl l + rewrite VecCast.cast-is-id ≡.refl p = ≡.refl Edgeₛ : (v : ℕ) → Setoid 0ℓ 0ℓ -Edgeₛ v = record { isEquivalence = ≈-Edge-IsEquivalence {v} } +Edgeₛ v = record { isEquivalence = ≈-IsEquivalence {v} } -- cgit v1.2.3