Skip to content

Commit

Permalink
Merge pull request #446 from jacquescomeaux/add-pushout-predicate
Browse files Browse the repository at this point in the history
Add predicate form for pushouts
  • Loading branch information
JacquesCarette authored Dec 10, 2024
2 parents f274d80 + 584abc0 commit 46bfebd
Show file tree
Hide file tree
Showing 2 changed files with 20 additions and 16 deletions.
16 changes: 9 additions & 7 deletions src/Categories/Diagram/Duality.agda
Original file line number Diff line number Diff line change
Expand Up @@ -76,13 +76,15 @@ coEqualizer⇒Coequalizer e = record

coPullback⇒Pushout : Pullback f g Pushout f g
coPullback⇒Pushout p = record
{ i₁ = p₁
; i₂ = p₂
; commute = commute
; universal = universal
; universal∘i₁≈h₁ = p₁∘universal≈h₁
; universal∘i₂≈h₂ = p₂∘universal≈h₂
; unique-diagram = unique-diagram
{ i₁ = p₁
; i₂ = p₂
; isPushout = record
{ commute = commute
; universal = universal
; universal∘i₁≈h₁ = p₁∘universal≈h₁
; universal∘i₂≈h₂ = p₂∘universal≈h₂
; unique-diagram = unique-diagram
}
}
where open Pullback p

Expand Down
20 changes: 11 additions & 9 deletions src/Categories/Diagram/Pushout.agda
Original file line number Diff line number Diff line change
Expand Up @@ -6,10 +6,6 @@ module Categories.Diagram.Pushout {o ℓ e} (C : Category o ℓ e) where

open Category C
open HomReasoning
open Equiv

open import Categories.Morphism.Reasoning C as Square
renaming (glue to glue-square) hiding (id-unique)

open import Level

Expand All @@ -18,11 +14,7 @@ private
A B E X Y Z : Obj
f g h j : A ⇒ B

record Pushout (f : X ⇒ Y) (g : X ⇒ Z) : Set (o ⊔ ℓ ⊔ e) where
field
{Q} : Obj
i₁ : Y ⇒ Q
i₂ : Z ⇒ Q
record IsPushout {Q : Obj} (f : X ⇒ Y) (g : X ⇒ Z) (i₁ : Y ⇒ Q) (i₂ : Z ⇒ Q) : Set (o ⊔ ℓ ⊔ e) where

field
commute : i₁ ∘ f ≈ i₂ ∘ g
Expand All @@ -39,3 +31,13 @@ record Pushout (f : X ⇒ Y) (g : X ⇒ Z) : Set (o ⊔ ℓ ⊔ e) where
unique-diagram
(j∘i₁≈h₁ ○ ⟺ universal∘i₁≈h₁)
(j∘i₂≈h₂ ○ ⟺ universal∘i₂≈h₂)

record Pushout (f : X ⇒ Y) (g : X ⇒ Z) : Set (o ⊔ ℓ ⊔ e) where

field
{Q} : Obj
i₁ : Y ⇒ Q
i₂ : Z ⇒ Q
isPushout : IsPushout f g i₁ i₂

open IsPushout isPushout public

0 comments on commit 46bfebd

Please sign in to comment.