Some graphs stuff (#98)

This commit is contained in:
Patrick Stevens
2020-02-15 13:41:51 +00:00
committed by GitHub
parent 9de323c5e8
commit f41f5226b9
18 changed files with 275 additions and 39 deletions

View File

@@ -8,6 +8,11 @@ open import Vectors
open import Semirings.Definition
open import Categories.Definition
open import Groups.Definition
open import Setoids.Setoids
open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Examples
open import Groups.Homomorphisms.Lemmas
open import Functions
module Categories.Examples where
@@ -20,13 +25,13 @@ Category.rightId (SET {a}) = λ f → refl
Category.leftId (SET {a}) = λ f refl
Category.compositionAssociative (SET {a}) = λ f g h refl
GROUP : {a b : _} Category {lsuc a b} {a b}
Category.objects (GROUP {a}) = Group {!!} {!!}
Category.arrows (GROUP {a}) = {!!}
Category.id (GROUP {a}) = {!!}
Category._∘_ (GROUP {a}) = {!!}
Category.rightId (GROUP {a}) = {!!}
Category.leftId (GROUP {a}) = {!!}
GROUP : {a b : _} Category {lsuc a lsuc b} {a b}
Category.objects (GROUP {a} {b}) = Sg (Set a) (λ A Sg (A A A) (λ _+_ Sg (Setoid {a} {b} A) (λ S Group S _+_)))
Category.arrows (GROUP {a}) (A , (_+A_ , (S , G))) (B , (_+B_ , (T , H))) = Sg (A B) (GroupHom G H)
Category.id (GROUP {a}) (A , (_+A_ , (S , G))) = (λ i i) , identityHom G
Category._∘_ (GROUP {a}) {A , (_+A_ , (S , G))} {B , (_+B_ , (T , H))} {C , (_+C_ , (U , I))} (f , fHom) (g , gHom) = (f g) , groupHomsCompose gHom fHom
Category.rightId (GROUP {a}) {A , (_+A_ , (S , G))} {B , (_+B_ , (T , H))} (f , fHom) = {!!}
Category.leftId (GROUP {a}) {A , (_+A_ , (S , G))} {B , (_+B_ , (T , H))} (f , fHom) = {!!}
Category.compositionAssociative (GROUP {a}) = {!!}
DISCRETE : {a : _} (X : Set a) Category {a} {a}
@@ -34,6 +39,6 @@ Category.objects (DISCRETE X) = X
Category.arrows (DISCRETE X) = λ a b a b
Category.id (DISCRETE X) = λ x refl
Category._∘_ (DISCRETE X) = λ y=z x=y transitivity x=y y=z
Category.rightId (DISCRETE X) = {!!}
Category.leftId (DISCRETE X) = {!!}
Category.compositionAssociative (DISCRETE X) = {!!}
Category.rightId (DISCRETE X) refl = refl
Category.leftId (DISCRETE X) refl = refl
Category.compositionAssociative (DISCRETE X) refl refl refl = refl

View File

@@ -133,4 +133,10 @@ open import Modules.Examples
open import Modules.Lemmas
open import Modules.DirectSum
open import Graphs.PathGraph
open import Graphs.CycleGraph
open import Graphs.UnionGraph
open import Graphs.CompleteGraph
open import Graphs.Colouring
module Everything.Safe where

21
Graphs/Colouring.agda Normal file
View File

@@ -0,0 +1,21 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Functions
open import Setoids.Setoids
open import Setoids.Subset
open import Graphs.Definition
open import Sets.FinSet.Definition
open import Numbers.Naturals.Semiring
module Graphs.Colouring where
Colouring : (n : ) {a b c : _} {V' : Set a} {V : Setoid {a} {b} V'} (G : Graph c V) Set (a c)
Colouring n {V' = V'} G = {x y : V'} (Graph._<->_ G x y) FinSet n
inheritColouring : {n : } {a b c : _} {V' : Set a} {V : Setoid {a} {b} V'} {G : Graph c V} (k : Colouring n G) {e : _} {pred : V' Set e} {sub : subset V pred} {d : _} {_<->'_ : Rel {_} {d} (Sg V' pred)} (H : Subgraph G sub _<->'_) Colouring n (subgraphIsGraph H)
inheritColouring k H {v1 , _} {v2 , _} x-y = k (Subgraph.inherits H x-y)
Monochromatic : {a b c : Level} {V' : Set a} {V : Setoid {a} {b} V'} {G : Graph c V} {n : } (k : Colouring n G) {d e : _} {pred : V' Set d} {sub : subset V pred} {_<->'_ : Rel {_} {e} (Sg V' pred)} (H : Subgraph G sub _<->'_) Set (a d e)
Monochromatic {V' = V'} {n = n} k {pred = pred} {sub = sub} {_<->'_} H = Sg (FinSet n) (λ c {x : Sg V' pred} {y : Sg V' pred} (edge : x <->' y) (inheritColouring k H) edge c)

28
Graphs/CompleteGraph.agda Normal file
View File

@@ -0,0 +1,28 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Functions
open import Setoids.Setoids
open import Setoids.Subset
open import Graphs.Definition
open import Sets.FinSet.Definition
open import Numbers.Naturals.Semiring
open import Sets.EquivalenceRelations
module Graphs.CompleteGraph where
CompleteGraph : {a b : _} {V' : Set a} (V : Setoid {a} {b} V') Graph b V
Graph._<->_ (CompleteGraph V) x y = ((Setoid.__ V x y) False)
Graph.noSelfRelation (CompleteGraph V) x x!=x = x!=x (Equivalence.reflexive (Setoid.eq V))
Graph.symmetric (CompleteGraph V) x!=y y=x = x!=y ((Equivalence.symmetric (Setoid.eq V)) y=x)
Graph.wellDefined (CompleteGraph V) {x} {y} {r} {s} x=y r=s x!=r y=s = x!=r (transitive x=y (transitive y=s (symmetric r=s)))
where
open Setoid V
open Equivalence eq
Kn : (n : ) Graph _ (reflSetoid (FinSet n))
Kn n = CompleteGraph (reflSetoid (FinSet n))
triangle : Graph _ (reflSetoid (FinSet 3))
triangle = Kn 3

28
Graphs/CycleGraph.agda Normal file
View File

@@ -0,0 +1,28 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Functions
open import Setoids.Setoids
open import Setoids.Subset
open import Graphs.Definition
open import Sets.FinSet.Definition
open import Sets.FinSet.Lemmas
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Sets.EquivalenceRelations
open import Graphs.PathGraph
module Graphs.CycleGraph where
CycleGraph : (n : ) .(2 <N n) Graph _ (reflSetoid (FinSet n))
Graph._<->_ (CycleGraph n _) x y = ((toNat x succ (toNat y)) || (toNat y succ (toNat x))) || (((toNat x 0) && (toNat y n)) || ((toNat y 0) && (toNat x n)))
Graph.noSelfRelation (CycleGraph n _) y (inl (inl x)) = nNotSucc x
Graph.noSelfRelation (CycleGraph n _) y (inl (inr x)) = nNotSucc x
Graph.noSelfRelation (CycleGraph (succ n) _) y (inr (inl (fst ,, snd))) = naughtE (transitivity (equalityCommutative fst) snd)
Graph.noSelfRelation (CycleGraph (succ n) _) y (inr (inr (fst ,, snd))) = naughtE (transitivity (equalityCommutative fst) snd)
Graph.symmetric (CycleGraph n _) (inl (inl x)) = inl (inr x)
Graph.symmetric (CycleGraph n _) (inl (inr x)) = inl (inl x)
Graph.symmetric (CycleGraph n _) (inr (inl x)) = inr (inr x)
Graph.symmetric (CycleGraph n _) (inr (inr x)) = inr (inl x)
Graph.wellDefined (CycleGraph n _) refl refl i = i

34
Graphs/Definition.agda Normal file
View File

@@ -0,0 +1,34 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Functions
open import Setoids.Setoids
open import Setoids.Subset
module Graphs.Definition where
record Graph {a b : _} (c : _) {V' : Set a} (V : Setoid {a} {b} V') : Set (a b lsuc c) where
field
_<->_ : Rel {a} {c} V'
noSelfRelation : (x : V') x <-> x False
symmetric : {x y : V'} x <-> y y <-> x
wellDefined : {x y r s : V'} Setoid.__ V x y Setoid.__ V r s x <-> r y <-> s
record GraphIso {a b c d e m : _} {V1' : Set a} {V2' : Set b} {V1 : Setoid {a} {c} V1'} {V2 : Setoid {b} {d} V2'} (G : Graph e V1) (H : Graph m V2) (f : V1' V2') : Set (a b c d e m) where
field
bij : SetoidBijection V1 V2 f
respects : (x y : V1') Graph._<->_ G x y Graph._<->_ H (f x) (f y)
respects' : (x y : V1') Graph._<->_ H (f x) (f y) Graph._<->_ G x y
record Subgraph {a b c d e : _} {V' : Set a} {V : Setoid {a} {b} V'} (G : Graph c V) {pred : V' Set d} (sub : subset V pred) (_<->'_ : Rel {_} {e} (Sg V' pred)) : Set (a b c d e) where
field
inherits : {x y : Sg V' pred} (x <->' y) (Graph._<->_ G (underlying x) (underlying y))
symmetric : {x y : Sg V' pred} (x <->' y) (y <->' x)
wellDefined : {x y r s : Sg V' pred} Setoid.__ (subsetSetoid V sub) x y Setoid.__ (subsetSetoid V sub) r s x <->' r y <->' s
subgraphIsGraph : {a b c d e : _} {V' : Set a} {V : Setoid {a} {b} V'} {G : Graph c V} {pred : V' Set d} {sub : subset V pred} {rel : Rel {_} {e} (Sg V' pred)} (H : Subgraph G sub rel) Graph e (subsetSetoid V sub)
Graph._<->_ (subgraphIsGraph {rel = rel} H) = rel
Graph.noSelfRelation (subgraphIsGraph {G = G} H) (v , isSub) v=v = Graph.noSelfRelation G v (Subgraph.inherits H v=v)
Graph.symmetric (subgraphIsGraph H) v1=v2 = Subgraph.symmetric H v1=v2
Graph.wellDefined (subgraphIsGraph H) = Subgraph.wellDefined H

26
Graphs/PathGraph.agda Normal file
View File

@@ -0,0 +1,26 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Functions
open import Setoids.Setoids
open import Setoids.Subset
open import Graphs.Definition
open import Sets.FinSet.Definition
open import Sets.FinSet.Lemmas
open import Numbers.Naturals.Semiring
open import Sets.EquivalenceRelations
module Graphs.PathGraph where
nNotSucc : {n : } (n succ n) False
nNotSucc {zero} ()
nNotSucc {succ n} pr = nNotSucc (succInjective pr)
PathGraph : (n : ) Graph _ (reflSetoid (FinSet (succ n)))
Graph._<->_ (PathGraph n) x y = (toNat x succ (toNat y)) || (toNat y succ (toNat x))
Graph.noSelfRelation (PathGraph n) x (inl bad) = nNotSucc bad
Graph.noSelfRelation (PathGraph n) x (inr bad) = nNotSucc bad
Graph.symmetric (PathGraph n) (inl x) = inr x
Graph.symmetric (PathGraph n) (inr x) = inl x
Graph.wellDefined (PathGraph n) refl refl i = i

View File

@@ -0,0 +1,23 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Functions
open import Setoids.Setoids
open import Setoids.Subset
open import Graphs.Definition
open import Sets.FinSet.Definition
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Sets.EquivalenceRelations
open import Graphs.CompleteGraph
open import Graphs.Colouring
module Graphs.RamseyTriangle where
hasMonochromaticTriangle : {a : _} {n : } (G : Graph a (reflSetoid (FinSet n))) Set (lsuc lzero)
hasMonochromaticTriangle {n = n} G = Sg (FinSet n Set) λ pred (subset (reflSetoid (FinSet n)) pred) && {!!}
ramseyForTriangle : (k : ) Sg (λ N (n : ) (N <N n) (c : Colouring k (Kn n)) {!!})
ramseyForTriangle k = {!!}

29
Graphs/UnionGraph.agda Normal file
View File

@@ -0,0 +1,29 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Functions
open import Setoids.Setoids
open import Setoids.DirectSum
open import Setoids.Subset
open import Graphs.Definition
open import Sets.FinSet.Definition
open import Sets.FinSet.Lemmas
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Sets.EquivalenceRelations
open import Graphs.PathGraph
module Graphs.UnionGraph where
unionGraph : {a b c d e f : _} {V' : Set a} {V : Setoid {a} {b} V'} (G : Graph c V) {W' : Set d} {W : Setoid {d} {e} W'} (H : Graph f W) Graph (c f) (directSumSetoid V W)
Graph._<->_ (unionGraph {c = c} {f = f} G H) (inl v) (inl v2) = embedLevel {c} {f} (Graph._<->_ G v v2)
Graph._<->_ (unionGraph G H) (inl v) (inr w) = False'
Graph._<->_ (unionGraph G H) (inr w) (inl v) = False'
Graph._<->_ (unionGraph {c = c} {f = f} G H) (inr w) (inr w2) = embedLevel {f} {c} (Graph._<->_ H w w2)
Graph.noSelfRelation (unionGraph G H) (inl v) (v=v ,, _) = Graph.noSelfRelation G v v=v
Graph.noSelfRelation (unionGraph G H) (inr w) (w=w ,, _) = Graph.noSelfRelation H w w=w
Graph.symmetric (unionGraph G H) {inl x} {inl y} (x=y ,, _) = Graph.symmetric G x=y ,, record {}
Graph.symmetric (unionGraph G H) {inr x} {inr y} (x=y ,, _) = Graph.symmetric H x=y ,, record {}
Graph.wellDefined (unionGraph G H) {inl x} {inl y} {inl z} {inl w} (x=y ,, _) (y=z ,, _) (z=w ,, _) = Graph.wellDefined G x=y y=z z=w ,, record {}
Graph.wellDefined (unionGraph G H) {inr x} {inr y} {inr z} {inr w} (x=y ,, _) (y=z ,, _) (z=w ,, _) = Graph.wellDefined H x=y y=z z=w ,, record {}

View File

@@ -7,7 +7,7 @@ open import Groups.Abelian.Definition
module Groups.Abelian.DirectSum {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+1_ : A A A} {_+2_ : B B B} {G1' : Group S _+1_} {G2' : Group T _+2_} (G1 : AbelianGroup G1') (G2 : AbelianGroup G2') where
open import Groups.DirectSum.Definition G1' G2'
open import Setoids.DirectSum S T
open import Setoids.Product S T
directSumAbGroup : AbelianGroup directSumGroup
AbelianGroup.commutative directSumAbGroup = directSumLift (AbelianGroup.commutative G1) (AbelianGroup.commutative G2)
AbelianGroup.commutative directSumAbGroup = productLift (AbelianGroup.commutative G1) (AbelianGroup.commutative G2)

View File

@@ -6,9 +6,9 @@ open import Groups.Definition
module Groups.DirectSum.Definition {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A A A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B B B} (G : Group S _·A_) (H : Group T _·B_) where
open import Setoids.DirectSum S T
open import Setoids.Product S T
directSumGroup : Group directSumSetoid (λ x1 y1 (((_&&_.fst x1) ·A (_&&_.fst y1)) ,, ((_&&_.snd x1) ·B (_&&_.snd y1))))
directSumGroup : Group productSetoid (λ x1 y1 (((_&&_.fst x1) ·A (_&&_.fst y1)) ,, ((_&&_.snd x1) ·B (_&&_.snd y1))))
Group.+WellDefined directSumGroup (s ,, t) (u ,, v) = Group.+WellDefined G s u ,, Group.+WellDefined H t v
Group.0G directSumGroup = (Group.0G G ,, Group.0G H)
Group.inverse directSumGroup (g1 ,, H1) = (Group.inverse G g1) ,, (Group.inverse H H1)

View File

@@ -38,3 +38,11 @@ listZip : {a b c : _} {A : Set a} {B : Set b} {C : Set c} (f : A → B → C)
listZip f f1 f2 [] l2 = map f2 l2
listZip f f1 f2 (x :: l1) [] = map f1 (x :: l1)
listZip f f1 f2 (x :: l1) (y :: l2) = (f x y) :: listZip f f1 f2 l1 l2
contains : {a : _} {A : Set a} (l : List A) (needle : A) Set a
contains [] needle = False'
contains (x :: l) needle = (x needle) || contains l needle
containsMap : {a b : _} {A : Set a} {B : Set b} (f : A B) (l : List A) (needle : A) (contains l needle) contains (map f l) (f needle)
containsMap f (x :: l) needle (inl cont) = inl (applyEquality f cont)
containsMap f (x :: l) needle (inr cont) = inr (containsMap f l needle cont)

View File

@@ -116,3 +116,6 @@ decidableOr : {a b : _} → (A : Set a) → (B : Set b) → (A || (A → False))
decidableOr {a} {b} A B decidable (inl x) = inl x
decidableOr {a} {b} A B (inl y) (inr x) = inl y
decidableOr {a} {b} A B (inr y) (inr x) = inr (record { fst = y ; snd = x})
embedLevel : {a b : _} Set a Set (a b)
embedLevel {a} {b} A = A && (True' {b})

View File

@@ -11,11 +11,11 @@ open import Modules.Definition
module Modules.DirectSum {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+R_ : A A A} {_*R_ : A A A} (R : Ring S _+R_ _*R_) {m n o p : _} {M : Set m} {T : Setoid {m} {n} M} {_+_ : M M M} {G' : Group T _+_} {G : AbelianGroup G'} {_·1_ : A M M} {N : Set o} {U : Setoid {o} {p} N} {_+'_ : N N N} {H' : Group U _+'_} {H : AbelianGroup H'} {_·2_ : A N N} (M1 : Module R G _·1_) (M2 : Module R H _·2_) where
open import Groups.Abelian.DirectSum G H
open import Setoids.DirectSum T U
open import Setoids.Product T U
directSumModule : Module R directSumAbGroup λ r mn ((r ·1 (_&&_.fst mn)) ,, (r ·2 (_&&_.snd mn)))
Module.dotWellDefined directSumModule r=s t=u = directSumLift (Module.dotWellDefined M1 r=s (_&&_.fst t=u)) (Module.dotWellDefined M2 r=s (_&&_.snd t=u))
Module.dotDistributesLeft directSumModule = directSumLift (Module.dotDistributesLeft M1) (Module.dotDistributesLeft M2)
Module.dotDistributesRight directSumModule = directSumLift (Module.dotDistributesRight M1) (Module.dotDistributesRight M2)
Module.dotAssociative directSumModule = directSumLift (Module.dotAssociative M1) (Module.dotAssociative M2)
Module.dotIdentity directSumModule = directSumLift (Module.dotIdentity M1) (Module.dotIdentity M2)
Module.dotWellDefined directSumModule r=s t=u = productLift (Module.dotWellDefined M1 r=s (_&&_.fst t=u)) (Module.dotWellDefined M2 r=s (_&&_.snd t=u))
Module.dotDistributesLeft directSumModule = productLift (Module.dotDistributesLeft M1) (Module.dotDistributesLeft M2)
Module.dotDistributesRight directSumModule = productLift (Module.dotDistributesRight M1) (Module.dotDistributesRight M2)
Module.dotAssociative directSumModule = productLift (Module.dotAssociative M1) (Module.dotAssociative M2)
Module.dotIdentity directSumModule = productLift (Module.dotIdentity M1) (Module.dotIdentity M2)

View File

@@ -29,3 +29,8 @@ aIsNotSuccA (succ a) pr = aIsNotSuccA a (succInjective pr)
DecideEquality (succ a) (succ b) with DecideEquality a b
DecideEquality (succ a) (succ b) | inl x = inl (applyEquality succ x)
DecideEquality (succ a) (succ b) | inr x = inr λ pr x (succInjective pr)
DecideEquality' : (a b : ) Bool
DecideEquality' a b with DecideEquality a b
DecideEquality' a b | inl x = BoolTrue
DecideEquality' a b | inr x = BoolFalse

View File

@@ -8,7 +8,7 @@ open import Rings.Definition
module Rings.DirectSum {a b c d : _} {A : Set a} {S : Setoid {a} {b} A} {_+1_ : A A A} {_*1_ : A A A} {C : Set c} {T : Setoid {c} {d} C} {_+2_ : C C C} {_*2_ : C C C} (R1 : Ring S _+1_ _*1_) (R2 : Ring T _+2_ _*2_) where
open import Setoids.DirectSum S T
open import Setoids.Product S T
pairPlus : A && C A && C A && C
pairPlus (a ,, b) (c ,, d) = (a +1 c) ,, (b +2 d)
@@ -16,12 +16,12 @@ pairPlus (a ,, b) (c ,, d) = (a +1 c) ,, (b +2 d)
pairTimes : A && C A && C A && C
pairTimes (a ,, b) (c ,, d) = (a *1 c) ,, (b *2 d)
directSumRing : Ring directSumSetoid pairPlus pairTimes
directSumRing : Ring productSetoid pairPlus pairTimes
Ring.additiveGroup directSumRing = directSumGroup (Ring.additiveGroup R1) (Ring.additiveGroup R2)
Ring.*WellDefined directSumRing r=t s=u = directSumLift (Ring.*WellDefined R1 (_&&_.fst r=t) (_&&_.fst s=u)) (Ring.*WellDefined R2 (_&&_.snd r=t) (_&&_.snd s=u))
Ring.*WellDefined directSumRing r=t s=u = productLift (Ring.*WellDefined R1 (_&&_.fst r=t) (_&&_.fst s=u)) (Ring.*WellDefined R2 (_&&_.snd r=t) (_&&_.snd s=u))
Ring.1R directSumRing = Ring.1R R1 ,, Ring.1R R2
Ring.groupIsAbelian directSumRing = directSumLift (Ring.groupIsAbelian R1) (Ring.groupIsAbelian R2)
Ring.*Associative directSumRing = directSumLift (Ring.*Associative R1) (Ring.*Associative R2)
Ring.*Commutative directSumRing = directSumLift (Ring.*Commutative R1) (Ring.*Commutative R2)
Ring.*DistributesOver+ directSumRing = directSumLift (Ring.*DistributesOver+ R1) (Ring.*DistributesOver+ R2)
Ring.identIsIdent directSumRing = directSumLift (Ring.identIsIdent R1) (Ring.identIsIdent R2)
Ring.groupIsAbelian directSumRing = productLift (Ring.groupIsAbelian R1) (Ring.groupIsAbelian R2)
Ring.*Associative directSumRing = productLift (Ring.*Associative R1) (Ring.*Associative R2)
Ring.*Commutative directSumRing = productLift (Ring.*Commutative R1) (Ring.*Commutative R2)
Ring.*DistributesOver+ directSumRing = productLift (Ring.*DistributesOver+ R1) (Ring.*DistributesOver+ R2)
Ring.identIsIdent directSumRing = productLift (Ring.identIsIdent R1) (Ring.identIsIdent R2)

View File

@@ -1,19 +1,20 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Sets.EquivalenceRelations
open import Setoids.Setoids
module Setoids.DirectSum {m n o p : _} {A : Set m} {B : Set n} (R : Setoid {m} {o} A) (S : Setoid {n} {p} B) where
open Setoid
directSumSetoid : Setoid (A && B)
__ (directSumSetoid) (a ,, b) (c ,, d) = (Setoid.__ R a c) && (Setoid.__ S b d)
Equivalence.reflexive (eq (directSumSetoid)) {(a ,, b)} = Equivalence.reflexive (Setoid.eq R) ,, Equivalence.reflexive (Setoid.eq S)
Equivalence.symmetric (eq (directSumSetoid)) {(a ,, b)} {(c ,, d)} (fst ,, snd) = Equivalence.symmetric (Setoid.eq R) fst ,, Equivalence.symmetric (Setoid.eq S) snd
Equivalence.transitive (eq (directSumSetoid)) {a ,, b} {c ,, d} {e ,, f} (fst1 ,, snd1) (fst2 ,, snd2) = Equivalence.transitive (Setoid.eq R) fst1 fst2 ,, Equivalence.transitive (Setoid.eq S) snd1 snd2
directSumLift : {r s : A} {t u : B} (Setoid.__ R r s) (Setoid.__ S t u) Setoid.__ directSumSetoid (r ,, t) (s ,, u)
_&&_.fst (directSumLift r=s t=u) = r=s
_&&_.snd (directSumLift r=s t=u) = t=u
directSumSetoid : Setoid {m n} (A || B)
Setoid.__ directSumSetoid (inl x) (inl y) = embedLevel {o} {p} (Setoid.__ R x y)
Setoid.__ directSumSetoid (inl x) (inr y) = False'
Setoid.__ directSumSetoid (inr x) (inl y) = False'
Setoid.__ directSumSetoid (inr x) (inr y) = embedLevel {p} {o} (Setoid.__ S x y)
Equivalence.reflexive (Setoid.eq directSumSetoid) {inl x} = Equivalence.reflexive (Setoid.eq R) {x} ,, record {}
Equivalence.reflexive (Setoid.eq directSumSetoid) {inr x} = Equivalence.reflexive (Setoid.eq S) {x} ,, record {}
Equivalence.symmetric (Setoid.eq directSumSetoid) {inl x} {inl y} (x=y ,, _) = (Equivalence.symmetric (Setoid.eq R) x=y) ,, record {}
Equivalence.symmetric (Setoid.eq directSumSetoid) {inr x} {inr y} (x=y ,, _) = (Equivalence.symmetric (Setoid.eq S) x=y) ,, record {}
Equivalence.transitive (Setoid.eq directSumSetoid) {inl x} {inl y} {inl z} (x=y ,, _) (y=z ,, _) = Equivalence.transitive (Setoid.eq R) x=y y=z ,, record {}
Equivalence.transitive (Setoid.eq directSumSetoid) {inr x} {inr y} {inr z} (x=y ,, _) (y=z ,, _) = Equivalence.transitive (Setoid.eq S) x=y y=z ,, record {}

19
Setoids/Product.agda Normal file
View File

@@ -0,0 +1,19 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Sets.EquivalenceRelations
open import Setoids.Setoids
module Setoids.Product {m n o p : _} {A : Set m} {B : Set n} (R : Setoid {m} {o} A) (S : Setoid {n} {p} B) where
open Setoid
productSetoid : Setoid (A && B)
__ (productSetoid) (a ,, b) (c ,, d) = (Setoid.__ R a c) && (Setoid.__ S b d)
Equivalence.reflexive (eq (productSetoid)) {(a ,, b)} = Equivalence.reflexive (Setoid.eq R) ,, Equivalence.reflexive (Setoid.eq S)
Equivalence.symmetric (eq (productSetoid)) {(a ,, b)} {(c ,, d)} (fst ,, snd) = Equivalence.symmetric (Setoid.eq R) fst ,, Equivalence.symmetric (Setoid.eq S) snd
Equivalence.transitive (eq (productSetoid)) {a ,, b} {c ,, d} {e ,, f} (fst1 ,, snd1) (fst2 ,, snd2) = Equivalence.transitive (Setoid.eq R) fst1 fst2 ,, Equivalence.transitive (Setoid.eq S) snd1 snd2
productLift : {r s : A} {t u : B} (Setoid.__ R r s) (Setoid.__ S t u) Setoid.__ productSetoid (r ,, t) (s ,, u)
_&&_.fst (productLift r=s t=u) = r=s
_&&_.snd (productLift r=s t=u) = t=u