Use different graph operations to implement construction

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
2024-04-25 23:10:41 -07:00
parent b134c143ca
commit c00c8e3e85
4 changed files with 239 additions and 197 deletions

View File

@@ -1,9 +1,11 @@
module Utils where
open import Agda.Primitive using () renaming (_⊔_ to _⊔_)
open import Data.Product as Prod using ()
open import Data.Nat using (; suc)
open import Data.List using (List; []; _∷_; _++_) renaming (map to mapˡ)
open import Data.List using (List; cartesianProduct; []; _∷_; _++_) renaming (map to mapˡ)
open import Data.List.Membership.Propositional using (_∈_)
open import Data.List.Membership.Propositional.Properties as ListMemProp using ()
open import Data.List.Relation.Unary.All using (All; []; _∷_; map)
open import Data.List.Relation.Unary.Any using (Any; here; there) -- TODO: re-export these with nicer names from map
open import Function.Definitions using (Injective)
@@ -78,3 +80,9 @@ proj₁ (v , _) = v
proj₂ : {a p q} {A : Set a} {P : A Set p} {Q : A Set q} {a : A} (P Q) a Q a
proj₂ (_ , v) = v
∈-cartesianProduct : {a b} {A : Set a} {B : Set b}
{x : A} {xs : List A} {y : B} {ys : List B}
x xs y ys (x Prod., y) cartesianProduct xs ys
∈-cartesianProduct {x = x} (here refl) y∈ys = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (x Prod.,_) y∈ys)
∈-cartesianProduct {x = x} {xs = x' _} {ys = ys} (there x∈rest) y∈ys = ListMemProp.∈-++⁺ʳ (mapˡ (x' Prod.,_) ys) (∈-cartesianProduct x∈rest y∈ys)