{-# OPTIONS --without-K --safe #-}
open import Level
open import Categories.Category.Core
open import Categories.Category.Cartesian using (Cartesian)
open import Categories.Category.BinaryProducts using (BinaryProducts)
open import Categories.Category.Cocartesian using (Cocartesian)
open import Categories.Category.Distributive using (Distributive)
open import Categories.Category.Extensive using (Extensive)
open import Categories.Diagram.Pullback using (Pullback)
import Categories.Morphism as M
import Categories.Morphism.Reasoning as MR
import Categories.Object.Coproduct as CP
open CP using (Coproduct; IsCoproduct; IsCoproduct⇒Coproduct)
import Categories.Object.Duality as Duality
module Categories.Category.Extensive.Properties.Distributive {o ℓ e} (𝒞 : Category o ℓ e) where
open Category 𝒞
open Pullback using (p₁∘universal≈h₁)
open M 𝒞
open MR 𝒞
open HomReasoning
open Equiv
open Duality 𝒞
Extensive×Cartesian⇒Distributive : Extensive 𝒞 → Cartesian 𝒞 → Distributive 𝒞
Extensive×Cartesian⇒Distributive extensive cartesian = record
{ cartesian = cartesian
; cocartesian = cocartesian
; isIsoˡ = record { inv = distrib.to ; iso = distrib.iso }
}
where
open Extensive extensive
open Cocartesian cocartesian
open Cartesian cartesian using (products)
module BP = BinaryProducts products
open BP
module _ {A B C : Obj} where
pb : ∀ {D} (g : D ⇒ B + C) → Pullback 𝒞 (π₂ {A = A} {B = B + C}) g
pb g = record { p₁ = id ⁂ g ; p₂ = π₂ ; isPullback = record
{ commute = π₂∘⁂
; universal = λ {_} {h₁} {h₂} H → ⟨ π₁ ∘ h₁ , h₂ ⟩
; p₁∘universal≈h₁ = λ {X} {h₁} {h₂} {eq} → begin
(id ⁂ g) ∘ ⟨ π₁ ∘ h₁ , h₂ ⟩ ≈⟨ ⁂∘⟨⟩ ⟩
⟨ id ∘ π₁ ∘ h₁ , g ∘ h₂ ⟩ ≈⟨ ⟨⟩-congʳ identityˡ ⟩
⟨ π₁ ∘ h₁ , g ∘ h₂ ⟩ ≈˘⟨ ⟨⟩-congˡ eq ⟩
⟨ π₁ ∘ h₁ , π₂ ∘ h₁ ⟩ ≈⟨ g-η ⟩
h₁ ∎
; p₂∘universal≈h₂ = project₂
; unique-diagram = λ {X} {h₁} {h₂} eq₁ eq₂ → BP.unique′ (begin
π₁ ∘ h₁ ≈⟨ pushˡ (introˡ refl) ⟩
id ∘ π₁ ∘ h₁ ≈⟨ extendʳ π₁∘⁂ ⟨
π₁ ∘ (id ⁂ g) ∘ h₁ ≈⟨ refl⟩∘⟨ eq₁ ⟩
π₁ ∘ (id ⁂ g) ∘ h₂ ≈⟨ extendʳ π₁∘⁂ ⟩
id ∘ π₁ ∘ h₂ ≈⟨ pullˡ (elimˡ refl) ⟩
π₁ ∘ h₂ ∎) eq₂
} }
distrib : (A × B) + (A × C) ≅ A × (B + C)
distrib = CP.up-to-iso 𝒞
coproduct
(IsCoproduct⇒Coproduct 𝒞 (pullback-of-cp-is-cp' (pb i₁) (pb i₂)))
module distrib = _≅_ distrib