aboutsummaryrefslogtreecommitdiff
path: root/Data/Hypergraph/Base.agda
blob: a2157fb5ff73b36ef4550c034093b2a6d4d4e5c6 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
{-# OPTIONS --without-K --safe #-}

module Data.Hypergraph.Base where

open import Data.Castable using (IsCastable)
open import Data.Nat.Base using ()
open import Data.Fin using (Fin)

import Data.Vec.Base as VecBase
import Data.Vec.Relation.Binary.Equality.Cast as VecCast
import Relation.Binary.PropositionalEquality as open import Relation.Binary using (Rel; IsDecTotalOrder)
open import Data.Nat.Base using ()
open import Level using (Level; suc; _⊔_)

record HypergraphLabel { : Level} : Set (suc ) where
  field
    Label :   Set     isCastable : IsCastable Label
    _[_≈_] : (n : )  Rel (Label n)     _[_≤_] : (n : )  Rel (Label n)     decTotalOrder : (n : )  IsDecTotalOrder (n [_≈_]) (n [_≤_])
  open IsCastable isCastable public

module Edge (HL : HypergraphLabel) where

  open HypergraphLabel HL using (Label; cast)
  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

module HypergraphList (HL : HypergraphLabel) where

  open import Data.List.Base using (List)
  open Edge HL using (Edge)

  record Hypergraph (v : ) : Set where
    field
      edges : List (Edge v)