open import Category -- https://github.com/konn/category-agda open import Level open import Category.Sets module SetsCompleteness where open import cat-utility open import Relation.Binary.Core open import Function import Relation.Binary.PropositionalEquality -- Extensionality a b = {A : Set a} {B : A → Set b} {f g : (x : A) → B x} → (∀ x → f x ≡ g x) → f ≡ g → ( λ x → f x ≡ λ x → g x ) postulate extensionality : { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) → Relation.Binary.PropositionalEquality.Extensionality c₂ c₂ ≡cong = Relation.Binary.PropositionalEquality.cong lemma1 : { c₂ : Level } {a b : Obj (Sets { c₂})} {f g : Hom Sets a b} → Sets [ f ≈ g ] → (x : a ) → f x ≡ g x lemma1 refl x = refl record Σ {a} (A : Set a) (B : Set a) : Set a where constructor _,_ field proj₁ : A proj₂ : B open Σ public SetsProduct : { c₂ : Level} → CreateProduct ( Sets { c₂} ) SetsProduct { c₂ } = record { product = λ a b → Σ a b ; π1 = λ a b → λ ab → (proj₁ ab) ; π2 = λ a b → λ ab → (proj₂ ab) ; isProduct = λ a b → record { _×_ = λ f g x → record { proj₁ = f x ; proj₂ = g x } -- ( f x , g x ) ; π1fxg=f = refl ; π2fxg=g = refl ; uniqueness = refl ; ×-cong = λ {c} {f} {f'} {g} {g'} f=f g=g → prod-cong a b f=f g=g } } where prod-cong : ( a b : Obj (Sets {c₂}) ) {c : Obj (Sets {c₂}) } {f f' : Hom Sets c a } {g g' : Hom Sets c b } → Sets [ f ≈ f' ] → Sets [ g ≈ g' ] → Sets [ (λ x → f x , g x) ≈ (λ x → f' x , g' x) ] prod-cong a b {c} {f} {.f} {g} {.g} refl refl = refl record iproduct {a} (I : Set a) ( pi0 : I → Set a ) : Set a where field pi1 : ( i : I ) → pi0 i open iproduct SetsIProduct : { c₂ : Level} → (I : Obj Sets) (ai : I → Obj Sets ) → IProduct ( Sets { c₂} ) I SetsIProduct I fi = record { ai = fi ; iprod = iproduct I fi ; pi = λ i prod → pi1 prod i ; isIProduct = record { iproduct = iproduct1 ; pif=q = pif=q ; ip-uniqueness = ip-uniqueness ; ip-cong = ip-cong } } where iproduct1 : {q : Obj Sets} → ((i : I) → Hom Sets q (fi i)) → Hom Sets q (iproduct I fi) iproduct1 {q} qi x = record { pi1 = λ i → (qi i) x } pif=q : {q : Obj Sets} (qi : (i : I) → Hom Sets q (fi i)) {i : I} → Sets [ Sets [ (λ prod → pi1 prod i) o iproduct1 qi ] ≈ qi i ] pif=q {q} qi {i} = refl ip-uniqueness : {q : Obj Sets} {h : Hom Sets q (iproduct I fi)} → Sets [ iproduct1 (λ i → Sets [ (λ prod → pi1 prod i) o h ]) ≈ h ] ip-uniqueness = refl ipcx : {q : Obj Sets} {qi qi' : (i : I) → Hom Sets q (fi i)} → ((i : I) → Sets [ qi i ≈ qi' i ]) → (x : q) → iproduct1 qi x ≡ iproduct1 qi' x ipcx {q} {qi} {qi'} qi=qi x = begin record { pi1 = λ i → (qi i) x } ≡⟨ ≡cong ( λ QIX → record { pi1 = QIX } ) ( extensionality Sets (λ i → ≡cong ( λ f → f x ) (qi=qi i) )) ⟩ record { pi1 = λ i → (qi' i) x } ∎ where open import Relation.Binary.PropositionalEquality open ≡-Reasoning ip-cong : {q : Obj Sets} {qi qi' : (i : I) → Hom Sets q (fi i)} → ((i : I) → Sets [ qi i ≈ qi' i ]) → Sets [ iproduct1 qi ≈ iproduct1 qi' ] ip-cong {q} {qi} {qi'} qi=qi = extensionality Sets ( ipcx qi=qi ) -- -- e f -- c -------→ a ---------→ b f ( f' -- ^ . ---------→ -- | . g -- |k . -- | . h --y : d -- cf. https://github.com/danr/Agda-projects/blob/master/Category-Theory/Equalizer.agda data sequ {c : Level} (A B : Set c) ( f g : A → B ) : Set c where elem : (x : A ) → (eq : f x ≡ g x) → sequ A B f g open sequ SetsEqualizer : { c₂ : Level} → (a b : Obj (Sets {c₂}) ) (f g : Hom (Sets {c₂}) a b) → Equalizer Sets f g SetsEqualizer {c₂} a b f g = record { equalizer-c = sequ a b f g ; equalizer = λ e → equ e ; isEqualizer = record { fe=ge = fe=ge ; k = k ; ek=h = λ {d} {h} {eq} → ek=h {d} {h} {eq} ; uniqueness = uniqueness } } where equ : ( sequ a b f g ) → a equ (elem x eq) = x fe=ge0 : (x : sequ a b f g) → (Sets [ f o (λ e → equ e) ]) x ≡ (Sets [ g o (λ e → equ e) ]) x fe=ge0 (elem x eq ) = eq fe=ge : Sets [ Sets [ f o (λ e → equ e ) ] ≈ Sets [ g o (λ e → equ e ) ] ] fe=ge = extensionality Sets (fe=ge0 ) k : {d : Obj Sets} (h : Hom Sets d a) → Sets [ Sets [ f o h ] ≈ Sets [ g o h ] ] → Hom Sets d (sequ a b f g) k {d} h eq = λ x → elem (h x) ( ≡cong ( λ y → y x ) eq ) ek=h : {d : Obj Sets} {h : Hom Sets d a} {eq : Sets [ Sets [ f o h ] ≈ Sets [ g o h ] ]} → Sets [ Sets [ (λ e → equ e ) o k h eq ] ≈ h ] ek=h {d} {h} {eq} = refl irr : {d : Set c₂ } { x y : d } ( eq eq' : x ≡ y ) → eq ≡ eq' irr refl refl = refl injection : { c₂ : Level } {a b : Obj (Sets { c₂})} (f : Hom Sets a b) → Set c₂ injection f = ∀ x y → f x ≡ f y → x ≡ y elm-cong : (x y : sequ a b f g) → equ x ≡ equ y → x ≡ y elm-cong ( elem x eq ) (elem .x eq' ) refl = ≡cong ( λ ee → elem x ee ) ( irr eq eq' ) lemma5 : {d : Obj Sets} {h : Hom Sets d a} {fh=gh : Sets [ Sets [ f o h ] ≈ Sets [ g o h ] ]} {k' : Hom Sets d (sequ a b f g)} → Sets [ Sets [ (λ e → equ e) o k' ] ≈ h ] → (x : d ) → equ (k h fh=gh x) ≡ equ (k' x) lemma5 refl x = refl -- somehow this is not equal to lemma1 uniqueness : {d : Obj Sets} {h : Hom Sets d a} {fh=gh : Sets [ Sets [ f o h ] ≈ Sets [ g o h ] ]} {k' : Hom Sets d (sequ a b f g)} → Sets [ Sets [ (λ e → equ e) o k' ] ≈ h ] → Sets [ k h fh=gh ≈ k' ] uniqueness {d} {h} {fh=gh} {k'} ek'=h = extensionality Sets ( λ ( x : d ) → begin k h fh=gh x ≡⟨ elm-cong ( k h fh=gh x) ( k' x ) (lemma5 {d} {h} {fh=gh} {k'} ek'=h x ) ⟩ k' x ∎ ) where open import Relation.Binary.PropositionalEquality open ≡-Reasoning open Functor record Small { c₁ c₂ ℓ : Level} ( C : Category c₁ c₂ ℓ ) ( I : Set c₁ ) : Set (suc (c₁ ⊔ c₂ ⊔ ℓ )) where field small→ : Obj C → I small← : I → Obj C small-iso : { x : Obj C } → Hom C (small← ( small→ x )) x shom→ : {i j : Obj C } → Hom C i j → I → I shom← : {i j : I } → ( f : I → I ) → Hom C ( small← i ) ( small← j ) shom-iso : {i j : I } → ( f : Hom C ( small← i ) ( small← j ) ) → C [ shom← ( shom→ f ) ≈ f ] open Small ΓObj : { c₁ c₂ ℓ : Level} { C : Category c₁ c₂ ℓ } { I : Set c₁ } ( s : Small C I ) ( Γ : Functor C ( Sets { c₁} )) (i : I ) →  Set c₁ ΓObj s Γ i = FObj Γ (small← s i ) ΓMap : { c₁ c₂ ℓ : Level} { C : Category c₁ c₂ ℓ } { I : Set c₁ } ( s : Small C I ) ( Γ : Functor C ( Sets { c₁} )) {i j : I } →  ( f : I → I ) → ΓObj s Γ i → ΓObj s Γ j ΓMap s Γ {i} {j} f = FMap Γ ( shom← s f ) record slim { c₂ } { I : Set c₂ } ( sobj : I → Set c₂ ) ( smap : { i j : I } → (f : I → I )→ sobj i → sobj j ) : Set c₂ where field slim-obj : ( i : I ) → sobj i open slim SetsLimit : { c₁ c₂ ℓ : Level} ( C : Category c₁ c₂ ℓ ) ( I : Set c₁ ) ( small : Small C I ) ( Γ : Functor C (Sets {c₁} ) ) → Limit Sets C Γ SetsLimit { c₂} C I s Γ = record { a0 = slim (ΓObj s Γ) (ΓMap s Γ) ; t0 = record { TMap = λ i → Sets [ proj i o e ] ; isNTrans = record { commute = comm1 } } ; isLimit = record { limit = {!!} ; t0f=t = {!!} ; limit-uniqueness = {!!} } } where a0 : Obj Sets a0 = slim (ΓObj s Γ) (ΓMap s Γ) iid : {i : Obj C } → Hom Sets (FObj Γ (small← s (small→ s i))) (FObj Γ i) iid {i} = FMap Γ ( small-iso s ) e : Hom Sets a0 (iproduct I (λ j → ΓObj s Γ j)) e = λ x → record { pi1 = λ j → slim-obj x j } proj : (i : Obj C ) → ( prod : iproduct I (λ j → ΓObj s Γ j )) → FObj Γ i proj i prod = iid ( pi1 prod ( small→ s i ) ) comm2 : {a b : Obj C} {f : Hom C a b} → ( x : a0 ) → (Sets [ FMap Γ f o Sets [ proj a o e ] ]) x ≡ (Sets [ proj b o e ]) x comm2 {a} {b} {f} x = begin (FMap Γ f ) ( ( proj a o e ) x ) ≡⟨⟩ (FMap Γ f ) ( iid ( slim-obj x (small→ s a) )) ≡⟨ {!!} ⟩ iid ( slim-obj x ( small→ s b ) ) ∎ where open import Relation.Binary.PropositionalEquality open ≡-Reasoning comm1 : {a b : Obj C} {f : Hom C a b} → Sets [ Sets [ FMap Γ f o Sets [ proj a o e ] ] ≈ Sets [ Sets [ proj b o e ] o FMap (K Sets C (slim (ΓObj s Γ) (ΓMap s Γ))) f ] ] comm1 {a} {b} {f} = begin Sets [ FMap Γ f o Sets [ proj a o e ] ] ≈⟨ extensionality Sets ( λ x → comm2 x ) ⟩ Sets [ proj b o e ] ≈↑⟨ idR ⟩ Sets [ Sets [ proj b o e ] o id1 Sets a0 ] ≈⟨⟩ Sets [ Sets [ proj b o e ] o FMap (K Sets C (slim (ΓObj s Γ) (ΓMap s Γ))) f ] ∎ where open import HomReasoning open ≈-Reasoning Sets