├── .gitignore ├── lean-toolchain ├── Topos.lean ├── .github └── workflows │ └── lean_action_ci.yml ├── lakefile.toml ├── README.md ├── lake-manifest.json └── Topos ├── Colimits.lean ├── Types.lean ├── Basic.lean ├── Power.lean ├── Classifier.lean └── Exponentials.lean /.gitignore: -------------------------------------------------------------------------------- 1 | /.lake 2 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.15.0 2 | -------------------------------------------------------------------------------- /Topos.lean: -------------------------------------------------------------------------------- 1 | -- This module serves as the root of the `Topos` library. 2 | -- Import modules here that should be built as part of the library. 3 | import Topos.Basic 4 | import Topos.Exponentials 5 | -------------------------------------------------------------------------------- /.github/workflows/lean_action_ci.yml: -------------------------------------------------------------------------------- 1 | name: Lean Action CI 2 | 3 | on: 4 | push: 5 | pull_request: 6 | workflow_dispatch: 7 | 8 | jobs: 9 | build: 10 | runs-on: ubuntu-latest 11 | 12 | steps: 13 | - uses: actions/checkout@v4 14 | - uses: leanprover/lean-action@v1 15 | -------------------------------------------------------------------------------- /lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "topos" 2 | version = "0.1.0" 3 | keywords = ["math"] 4 | defaultTargets = ["Topos"] 5 | 6 | [leanOptions] 7 | pp.unicode.fun = true # pretty-prints `fun a ↦ b` 8 | autoImplicit = false 9 | 10 | [[require]] 11 | name = "mathlib" 12 | scope = "leanprover-community" 13 | rev = "v4.15.0" 14 | 15 | [[lean_lib]] 16 | name = "Topos" 17 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # topos 2 | 3 | A lean 4 port of content from the lean 3 topos library at [b-mehta/topos](https://github.com/b-mehta/topos). The content of this repository follows Mac Lane and Moerdijk's "Sheaves in Geometry and Logic" and is currently incomplete. The eventual goal is to contribute this code to [mathlib4](https://github.com/leanprover-community/mathlib4). 4 | 5 | ## TODO 6 | 7 | - Finish `Topos/Colimits.lean` thereby verifying that a topos has all finite colimits. 8 | - Create `Topos/Example.lean` and formally verify that existing instances of topoi (sets, sheaves on sites) are topoi. 9 | - Chores relating to making this repository [mathlib friendly](https://leanprover-community.github.io/contribute/index.html). 10 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover-community/mathlib4", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "leanprover-community", 8 | "rev": "9837ca9d65d9de6fad1ef4381750ca688774e608", 9 | "name": "mathlib", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "v4.15.0", 12 | "inherited": false, 13 | "configFile": "lakefile.lean"}, 14 | {"url": "https://github.com/leanprover-community/plausible", 15 | "type": "git", 16 | "subDir": null, 17 | "scope": "leanprover-community", 18 | "rev": "2c57364ef83406ea86d0f78ce3e342079a2fece5", 19 | "name": "plausible", 20 | "manifestFile": "lake-manifest.json", 21 | "inputRev": "v4.15.0", 22 | "inherited": true, 23 | "configFile": "lakefile.toml"}, 24 | {"url": "https://github.com/leanprover-community/LeanSearchClient", 25 | "type": "git", 26 | "subDir": null, 27 | "scope": "leanprover-community", 28 | "rev": "003ff459cdd85de551f4dcf95cdfeefe10f20531", 29 | "name": "LeanSearchClient", 30 | "manifestFile": "lake-manifest.json", 31 | "inputRev": "main", 32 | "inherited": true, 33 | "configFile": "lakefile.toml"}, 34 | {"url": "https://github.com/leanprover-community/import-graph", 35 | "type": "git", 36 | "subDir": null, 37 | "scope": "leanprover-community", 38 | "rev": "9a0b533c2fbd6195df067630be18e11e4349051c", 39 | "name": "importGraph", 40 | "manifestFile": "lake-manifest.json", 41 | "inputRev": "v4.15.0", 42 | "inherited": true, 43 | "configFile": "lakefile.toml"}, 44 | {"url": "https://github.com/leanprover-community/ProofWidgets4", 45 | "type": "git", 46 | "subDir": null, 47 | "scope": "leanprover-community", 48 | "rev": "2b000e02d50394af68cfb4770a291113d94801b5", 49 | "name": "proofwidgets", 50 | "manifestFile": "lake-manifest.json", 51 | "inputRev": "v0.0.48", 52 | "inherited": true, 53 | "configFile": "lakefile.lean"}, 54 | {"url": "https://github.com/leanprover-community/aesop", 55 | "type": "git", 56 | "subDir": null, 57 | "scope": "leanprover-community", 58 | "rev": "2689851f387bb2cef351e6825fe94a56a304ca13", 59 | "name": "aesop", 60 | "manifestFile": "lake-manifest.json", 61 | "inputRev": "v4.15.0", 62 | "inherited": true, 63 | "configFile": "lakefile.toml"}, 64 | {"url": "https://github.com/leanprover-community/quote4", 65 | "type": "git", 66 | "subDir": null, 67 | "scope": "leanprover-community", 68 | "rev": "f0c584bcb14c5adfb53079781eeea75b26ebbd32", 69 | "name": "Qq", 70 | "manifestFile": "lake-manifest.json", 71 | "inputRev": "v4.15.0", 72 | "inherited": true, 73 | "configFile": "lakefile.toml"}, 74 | {"url": "https://github.com/leanprover-community/batteries", 75 | "type": "git", 76 | "subDir": null, 77 | "scope": "leanprover-community", 78 | "rev": "e8dc5fc16c625fc4fe08f42d625523275ddbbb4b", 79 | "name": "batteries", 80 | "manifestFile": "lake-manifest.json", 81 | "inputRev": "v4.15.0", 82 | "inherited": true, 83 | "configFile": "lakefile.toml"}, 84 | {"url": "https://github.com/leanprover/lean4-cli", 85 | "type": "git", 86 | "subDir": null, 87 | "scope": "leanprover", 88 | "rev": "0c8ea32a15a4f74143e4e1e107ba2c412adb90fd", 89 | "name": "Cli", 90 | "manifestFile": "lake-manifest.json", 91 | "inputRev": "main", 92 | "inherited": true, 93 | "configFile": "lakefile.toml"}], 94 | "name": "topos", 95 | "lakeDir": ".lake"} 96 | -------------------------------------------------------------------------------- /Topos/Colimits.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Charlie Conneen. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Charlie Conneen 5 | -/ 6 | import Mathlib.CategoryTheory.Monad.Monadicity 7 | import Topos.Basic 8 | 9 | namespace CategoryTheory 10 | 11 | open Category Limits MonoClassifier Power Functor 12 | 13 | namespace Topos 14 | 15 | universe u v 16 | variable {C} 17 | variable [Category.{v, u} C] [IsTopos C] 18 | 19 | noncomputable section 20 | 21 | -- TODO: prove that `powFunctor C` preserves colimits of reflexive pairs. 22 | 23 | namespace BeckChevalley 24 | 25 | /- 26 | In this section, we follow Mac Lane and Moerdijk in defining the direct image ∃ₖ : PB' ⟶ PB 27 | of a monomorphism k : B' ⟶ B, then show that ∃ₖ ≫ Pow_map k = 𝟙 PB'. 28 | -/ 29 | 30 | variable {B B' : C} (k : B' ⟶ B) [Mono k] 31 | 32 | #check transpose (χ_ ((pullback.fst (f := in_ B') (g := t C)) ≫ prod.map k (𝟙 _))) 33 | 34 | def directImage : pow B' ⟶ pow B := 35 | transpose (χ_ ((pullback.fst (f := in_ B') (g := t C)) ≫ prod.map k (𝟙 _))) 36 | 37 | variable {S : C} (m : S ⟶ B') [Mono m] 38 | 39 | lemma wDef_comm' : (prod.map m (𝟙 _)) ≫ (prod.map (𝟙 _) (name (χ_ m))) ≫ in_ B' = terminal.from _ ≫ t C := by 40 | rw [Predicate.NameDef, prod.map_fst_assoc] 41 | have h : terminal.from (S ⨯ ⊤_ C) = prod.fst ≫ terminal.from S := by apply terminal.hom_ext 42 | rw [h, assoc, MonoClassifier.comm] 43 | 44 | lemma wDef_comm : (prod.map m (name (χ_ m))) ≫ in_ B' = terminal.from _ ≫ t C := by 45 | -- for some reason there is an issue rewriting m = m ≫ 𝟙 _ ?? 46 | -- TODO: should be able to wrestle this lemma's statement into the previous lemma's, merging the two 47 | have h := wDef_comm' m 48 | rw [prod.map_map_assoc, comp_id, id_comp] at h 49 | assumption 50 | 51 | def w : S ⨯ ⊤_ C ⟶ pullback (in_ B') (t C) := pullback.lift (w := wDef_comm m) 52 | 53 | lemma directImage_NameChar_factors : name (χ_ m) ≫ directImage k = name (χ_ (m ≫ k)) := by 54 | have transpose : transposeInv (name (χ_ m) ≫ directImage k) = transposeInv (name (χ_ (m ≫ k))) := by 55 | dsimp only [name] 56 | rw [transpose_left_inv] 57 | dsimp only [transposeInv, directImage] 58 | rw [prod.map_id_comp, assoc, Power.comm] 59 | sorry 60 | 61 | sorry 62 | 63 | end BeckChevalley 64 | 65 | instance PowRightAdj : IsRightAdjoint (powFunctor C) where 66 | exists_leftAdjoint := by 67 | apply Exists.intro (powFunctorOp C) 68 | exact ⟨powSelfAdj C⟩ 69 | 70 | instance PowFaithful : Functor.Faithful (powFunctor C) where 71 | map_injective := by 72 | intro ⟨X⟩ ⟨Y⟩ ⟨f⟩ ⟨g⟩ h 73 | change (inverseImage f = inverseImage g) at h 74 | congr 75 | have h' := congrArg (fun k ↦ transposeInv (singleton X ≫ k)) h 76 | dsimp only [transposeInv] at h' 77 | rw [prod.map_id_comp, prod.map_id_comp, Category.assoc, Category.assoc, inverseImage_comm, inverseImage_comm, 78 | ←Category.assoc, prod.map_map, ←Category.assoc, prod.map_map, id_comp, comp_id, id_comp, ←comp_id f, 79 | ←id_comp (singleton _), ←comp_id g, ←prod.map_map, ←prod.map_map, assoc, assoc, singleton, Power.comm] at h' 80 | have comm : (f ≫ terminal.from _) ≫ t C = prod.lift (𝟙 _) f ≫ prod.map f (𝟙 _) ≫ Predicate.eq _ := by 81 | rw [terminal.comp_from, ←assoc, prod.lift_map, comp_id, id_comp, Predicate.lift_eq, Predicate.true_] 82 | rw [terminal.comp_from, h', ←assoc, prod.lift_map, id_comp, comp_id] at comm 83 | exact (Predicate.eq_of_lift_eq comm.symm).symm 84 | 85 | 86 | instance hasCoreflexiveEqualizers : HasCoreflexiveEqualizers C := 87 | hasCoreflexiveEqualizers_of_hasEqualizers C 88 | 89 | instance : HasCoequalizers Cᵒᵖ := hasCoequalizers_opposite 90 | 91 | instance : HasReflexiveCoequalizers Cᵒᵖ := hasReflexiveCoequalizers_of_hasCoequalizers Cᵒᵖ 92 | 93 | instance PowReflectsIsomorphisms : Functor.ReflectsIsomorphisms (powFunctor C) := reflectsIsomorphismsOp (F := powFunctor C) 94 | 95 | instance PowPreservesCoproductOfReflexivePair : Monad.PreservesColimitOfIsReflexivePair (powFunctor C) where 96 | out := by 97 | intro ⟨A⟩ ⟨B⟩ ⟨f⟩ ⟨g⟩ h₀ 98 | change (B ⟶ A) at f; change (B ⟶ A) at g 99 | have h₁ := h₀.common_section' 100 | let s := h₁.choose 101 | have hs₁ := congrArg (fun k ↦ k.unop) h₁.choose_spec.1 102 | have hs₂ := congrArg (fun k ↦ k.unop) h₁.choose_spec.2 103 | change (f ≫ s.unop = 𝟙 _) at hs₁ 104 | change (g ≫ s.unop = 𝟙 _) at hs₂ 105 | refine PreservesColimit.mk ?_ 106 | intro ⟨pt, ι⟩ hc 107 | 108 | 109 | sorry 110 | 111 | instance powFunctorMonadic : MonadicRightAdjoint (powFunctor C) := sorry 112 | 113 | -- TODO: Use `powFunctorMonadic` to show that a topos has finite colimits. 114 | 115 | instance HasFiniteColimits : HasFiniteColimits C := sorry 116 | 117 | 118 | end 119 | end Topos 120 | end CategoryTheory 121 | -------------------------------------------------------------------------------- /Topos/Types.lean: -------------------------------------------------------------------------------- 1 | import Topos.Basic 2 | import Mathlib.Order.Basic 3 | import Mathlib.CategoryTheory.Sites.Types 4 | import Mathlib.CategoryTheory.Limits.Shapes.Types 5 | import Mathlib.CategoryTheory.Monoidal.Types.Basic 6 | 7 | namespace CategoryTheory 8 | 9 | open CategoryTheory Category Topos Opposite Limits Classifier Types Classical 10 | 11 | namespace Topos 12 | 13 | universe u 14 | 15 | #check ChosenFiniteProducts 16 | #synth HasTerminal (Type u) 17 | #synth HasPullbacks (Type u) 18 | #synth HasProducts (Type u) 19 | #synth Category (Type u) 20 | 21 | #check Classical.byCases 22 | 23 | #check PUnit.{u} 24 | #check ⊤_ (Type u) 25 | 26 | noncomputable section 27 | 28 | variable {A B : Type u} 29 | variable {p : Prop} (C : Type u) 30 | 31 | -- "the" subobject classifier for `Type u` 32 | inductive Ω : Type u where 33 | | T 34 | | F 35 | 36 | abbrev T := Ω.T 37 | abbrev F := Ω.F 38 | 39 | lemma Ω.em (z : Ω) : z = T ∨ z = F := by 40 | cases z with 41 | | T => left; rfl 42 | | F => right; rfl 43 | 44 | abbrev true_ : ⊤_ (Type u) ⟶ Ω := fun _ => T 45 | 46 | abbrev false_ : ⊤_ (Type u) ⟶ Ω := fun _ => F 47 | 48 | abbrev char (f : A ⟶ B) [Mono f] : B ⟶ Ω := 49 | fun b => if (∃ a : A, f a = b) then T else F 50 | 51 | variable (f : A ⟶ B) [mono : Mono f] 52 | 53 | lemma f_inj : Function.Injective f := (mono_iff_injective f).mp mono 54 | 55 | lemma char_true (b : B) (h : ∃ a : A, f a = b) : char f b = T := by 56 | simp; assumption 57 | 58 | lemma char_true_iff (b : B) : (∃ a : A, f a = b) ↔ char f b = T := by simp 59 | 60 | lemma char_true_iff_uniq (b : B) : (∃! a : A, f a = b) ↔ char f b = T := by 61 | simp 62 | apply Iff.intro 63 | · exact fun ⟨a, h₁, _⟩ => ⟨a, h₁⟩ 64 | · exact fun ⟨a, h⟩ => 65 | ⟨a, h, fun _ h₁ => ((mono_iff_injective f).mp mono) (h₁.trans h.symm)⟩ 66 | 67 | def charCommSq : CommSq f (terminal.from A) (char f) true_ where 68 | w := by funext a; simp 69 | 70 | def pullbackCone : PullbackCone (char f) true_ := 71 | PullbackCone.mk f (terminal.from A) (charCommSq f).w 72 | 73 | variable {f} 74 | 75 | def lift {X : Type u} (g : X ⟶ B) (hg : g ≫ char f = terminal.from X ≫ true_) : 76 | X ⟶ A := 77 | fun x => ((char_true_iff f (g x)).mpr 78 | (by show char f (g x) = T; rw [←types_comp_apply g (char f), hg, types_comp_apply])).choose 79 | 80 | lemma lift_fac {X : Type u} (g : X ⟶ B) (hg : g ≫ char f = terminal.from X ≫ true_) : 81 | lift g hg ≫ f = g := by 82 | funext x 83 | exact ((char_true_iff f (g x)).mpr 84 | (by show char f (g x) = T; rw [←types_comp_apply g (char f), hg, types_comp_apply])).choose_spec 85 | 86 | instance typesLimit : IsLimit (pullbackCone f) where 87 | lift := by 88 | intro s 89 | intro x 90 | have h' : (char f) ((PullbackCone.fst s) x) = T := by 91 | change (PullbackCone.fst s ≫ (char f)) x = T 92 | rw [PullbackCone.condition s, types_comp_apply] 93 | exact ((char_true_iff f (PullbackCone.fst s x)).mpr h').choose 94 | fac := by 95 | intro s j 96 | cases j with 97 | | none => 98 | simp [pullbackCone] 99 | funext x 100 | rw [(charCommSq f).w] 101 | simp [true_, PullbackCone.condition s] 102 | rw [←types_comp_apply (PullbackCone.fst s) (char f) x, (PullbackCone.condition s)] 103 | simp 104 | | some val => 105 | cases val 106 | · simp 107 | funext x 108 | have h' : (char f) ((PullbackCone.fst s) x) = T := by 109 | change (PullbackCone.fst s ≫ (char f)) x = T 110 | rw [PullbackCone.condition s, types_comp_apply] 111 | 112 | exact ((char_true_iff f (PullbackCone.fst s x)).mpr h').choose_spec 113 | · simp 114 | have ha : (PullbackCone.snd s) = terminal.from s.pt := 115 | terminal.hom_ext (PullbackCone.snd s) (terminal.from s.pt) 116 | have hb : (pullbackCone f).snd = terminal.from A := 117 | terminal.hom_ext (pullbackCone f).snd (terminal.from A) 118 | rw [ha, hb] 119 | apply terminal.comp_from 120 | uniq := by 121 | intro s j hj 122 | funext x 123 | have h' : (char f) ((PullbackCone.fst s) x) = T := by 124 | change (PullbackCone.fst s ≫ (char f)) x = T 125 | rw [PullbackCone.condition s, types_comp_apply] 126 | apply f_inj f 127 | have hj' := hj (some WalkingPair.left) 128 | simp [pullbackCone] at hj' 129 | rw [←types_comp_apply j f, hj'] 130 | exact ((char_true_iff f (PullbackCone.fst s x)).mpr h').choose_spec.symm 131 | 132 | 133 | def isPullback (f : A ⟶ B) [Mono f] : IsPullback f (terminal.from _) (char f) true_ where 134 | w := (charCommSq f).w 135 | isLimit' := Nonempty.intro typesLimit 136 | 137 | def classifier : Classifier (Type u) where 138 | t := true_ 139 | char := char 140 | isPullback := isPullback 141 | uniq := by 142 | intro A B f inst χ hχ 143 | funext x 144 | simp only 145 | if h : (∃ a : A, f a = x) then 146 | simp [char, h] 147 | rw [h.choose_spec.symm, ←types_comp_apply f χ, hχ.w] 148 | simp 149 | else 150 | simp [char, h] 151 | by_contra h' 152 | have h'' : χ x = T := by 153 | cases (Ω.em (χ x)) with 154 | | inl _ => assumption 155 | | inr _ => contradiction 156 | have k := hχ.w 157 | let A' := {b : B // (∃ a : A, f a = b) ∨ b = x} 158 | let f' : A' ⟶ B := fun b => (b : B) 159 | have hA' : f' ≫ χ = terminal.from A' ≫ true_ := by 160 | funext y 161 | simp [true_, f'] 162 | cases y.property with 163 | | inl k' => 164 | have r := congrArg χ k'.choose_spec 165 | rw [←types_comp_apply f χ k'.choose, k] at r 166 | simp [true_] at r 167 | exact r.symm 168 | | inr k' => rw [k']; assumption 169 | have lift := hχ.lift _ _ hA' 170 | have char_x : char f x = F := by 171 | dsimp [char] 172 | simp [h] 173 | have app : (hχ.lift _ _ hA' ≫ f ≫ char f) ⟨x, by simp⟩ = T := by simp 174 | have final : char f x = T := by 175 | rw [←assoc] at app 176 | rw [hχ.lift_fst _ _ hA', types_comp_apply] at app 177 | dsimp [f'] at app 178 | assumption 179 | have s : T = F := final.symm.trans char_x 180 | contradiction 181 | 182 | instance typesHasClassifier : HasClassifier (Type u) where 183 | exists_classifier := ⟨classifier⟩ 184 | 185 | variable (B) 186 | 187 | -- needs to use "a" subobject classifier. can't use just `Ω` 188 | def pow : Type u := (B ⟶ HasClassifier.Ω _) 189 | 190 | def in_ : B × (pow B) ⟶ HasClassifier.Ω _ := 191 | fun (b, f) => f b 192 | 193 | --def powerObject : Power.PowerObject B where 194 | -- pow := pow B 195 | -- in_ := in_ B 196 | 197 | instance typeIsTopos : IsTopos (Type u) where 198 | has_power_objects := sorry 199 | 200 | end 201 | end CategoryTheory.Topos 202 | -------------------------------------------------------------------------------- /Topos/Basic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Charlie Conneen. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Charlie Conneen 5 | -/ 6 | import Mathlib.CategoryTheory.ChosenFiniteProducts 7 | import Mathlib.CategoryTheory.Limits.Constructions.Equalizers 8 | -- import Mathlib.CategoryTheory.Topos.Power 9 | import Topos.Power 10 | 11 | 12 | /-! 13 | # Topoi 14 | 15 | We follow Mac Lane and Moerdijk in defining a topos as a category 16 | with a terminal object, pullbacks, a subobject classifier, and 17 | power objects. 18 | 19 | In this file, a "predicate" refers to any morphism into `Ω C`. 20 | 21 | ## Main definitions 22 | 23 | * `CategoryTheory.IsTopos C` is a typeclass asserting that `C` is a topos. 24 | 25 | * The namespace `CategoryTheory.Topos.Predicate` provides an API for various 26 | useful predicates, such as `Predicate.eq` as the equality predicate 27 | `B ⨯ B ⟶ Ω C`, and `Predicate.isSingleton` which is the classifier of the 28 | singleton map `singleton B : B ⟶ Pow B`. 29 | 30 | ## Main results 31 | 32 | * `singleton B : B ⟶ Pow B` is a monomorphism. This is `singletonMono`. 33 | 34 | ## Notation 35 | 36 | * If `φ` is a predicate `X ⟶ Ω C`, `⌈φ⌉` is shorthand for `Name φ`, 37 | which is the global element `⊤_ C ⟶ Pow X` associated to `φ`. 38 | 39 | ## References 40 | 41 | * [S. MacLane and I. Moerdijk, *Sheaves in Geometry and Logic*][MLM92] 42 | 43 | -/ 44 | 45 | namespace CategoryTheory 46 | 47 | open Category Limits HasClassifier Power 48 | 49 | universe u v 50 | 51 | variable (C : Type u) [Category.{v} C] [ChosenFiniteProducts C] 52 | 53 | /-- A category is a topos if it has a terminal object, 54 | all pullbacks, a subobject classifier, and power objects. 55 | -/ 56 | class IsTopos where 57 | /-- A topos has pullbacks. -/ 58 | [has_pullbacks : HasPullbacks C] 59 | /-- A topos has a subobject classifier. -/ 60 | [subobject_classifier : HasClassifier C] 61 | /-- A topos has power objects. -/ 62 | [has_power_objects : HasPowerObjects C] 63 | 64 | attribute [instance] IsTopos.has_pullbacks 65 | IsTopos.subobject_classifier 66 | IsTopos.has_power_objects 67 | 68 | variable [IsTopos C] {C} 69 | 70 | namespace Topos 71 | 72 | -- /-- A topos has (chosen) finite products. -/ 73 | -- noncomputable instance chosenFiniteProducts : ChosenFiniteProducts C := 74 | -- ChosenFiniteProducts.ofFiniteProducts C 75 | /-- A topos has binary products. -/ 76 | instance hasBinaryProducts : HasBinaryProducts C := 77 | hasBinaryProducts_of_hasTerminal_and_pullbacks C 78 | /-- A topos has finite products. -/ 79 | instance hasFiniteProducts : HasFiniteProducts C := 80 | hasFiniteProducts_of_has_binary_and_terminal 81 | /-- A topos has equalizers, since it has pullbacks and 82 | binary products. 83 | -/ 84 | instance hasEqualizers : HasEqualizers C := 85 | hasEqualizers_of_hasPullbacks_and_binary_products 86 | 87 | 88 | noncomputable section 89 | 90 | /-- The predicate on `B` which corresponds to the subobject `𝟙 B: B ⟶ B`. -/ 91 | abbrev Predicate.true_ (B : C) : B ⟶ Ω C := terminal.from B ≫ (t C) 92 | 93 | /-- 94 | The equality predicate on `B ⨯ B`. 95 | -/ 96 | abbrev Predicate.eq (B : C) : B ⨯ B ⟶ Ω C := χ_ (diag B) 97 | 98 | /-- The lift `X ⟶ B ⨯ B` of a morphism with itself, when composed 99 | with `predicate.eq B`, is true. 100 | -/ 101 | lemma Predicate.lift_eq {X B : C} (b : X ⟶ B) : 102 | prod.lift b b ≫ Predicate.eq B = Predicate.true_ X := by 103 | rw [←prod.comp_diag b, assoc, comm (diag B), ←assoc, terminal.comp_from] 104 | 105 | /-- Two maps in a topos are equal if their lift composed with 106 | the equality predicate on `B ⨯ B` is true. 107 | In other words, this combined with `Predicate.lift_eq` states that 108 | `Predicate.eq` is able to distinguish whether two morphisms are equal. 109 | -/ 110 | lemma Predicate.eq_of_lift_eq {X B : C} {b b' : X ⟶ B} 111 | (comm' : prod.lift b b' ≫ Predicate.eq B = Predicate.true_ X) : 112 | b = b' := by 113 | dsimp only [true_] at comm' 114 | have t : (IsPullback.lift (isPullback _) _ _ comm') ≫ diag _ = prod.lift b b' := 115 | IsPullback.lift_fst (isPullback _) _ _ comm' 116 | rw [prod.comp_diag] at t 117 | have t₁ := congrArg (fun k ↦ k ≫ prod.fst) t 118 | have t₂ := congrArg (fun k ↦ k ≫ prod.snd) t 119 | simp at t₁ 120 | simp at t₂ 121 | exact t₁.symm.trans t₂ 122 | 123 | /-- The "singleton" map `B ⟶ Pow B`. 124 | In Set, this map sends `b ∈ B` to the 125 | singleton set containing just `b`. 126 | -/ 127 | def singleton (B : C) : B ⟶ pow B := (Predicate.eq B)^ 128 | 129 | /-- `singleton B : B ⟶ Pow B` is a monomorphism. -/ 130 | instance singletonMono (B : C) : Mono (singleton B) where 131 | right_cancellation := by 132 | intro X b b' h 133 | rw [singleton] at h 134 | have h₁ : prod.map (𝟙 _) (b ≫ (Predicate.eq B)^) ≫ in_ B 135 | = prod.map (𝟙 _) (b' ≫ (Predicate.eq B)^) ≫ in_ B := 136 | congrFun (congrArg CategoryStruct.comp (congrArg (prod.map (𝟙 B)) h)) (in_ B) 137 | rw [prod.map_id_comp, assoc, Power.comm, prod.map_id_comp, assoc, Power.comm] at h₁ 138 | have comm : (b ≫ terminal.from _) ≫ t C 139 | = prod.lift b (𝟙 _) ≫ prod.map (𝟙 _) b ≫ Predicate.eq _ := by 140 | rw [terminal.comp_from, ←assoc, prod.lift_map, comp_id, 141 | id_comp, Predicate.lift_eq, Predicate.true_] 142 | rw [terminal.comp_from, h₁, ←assoc, prod.lift_map, id_comp, comp_id] at comm 143 | exact Predicate.eq_of_lift_eq comm.symm 144 | 145 | /-- The predicate on `Pow B` which picks out the subobject of "singletons". 146 | -/ 147 | def Predicate.isSingleton (B : C) : pow B ⟶ Ω C := χ_ (singleton B) 148 | 149 | /-- The name ⌈φ⌉ : ⊤_ C ⟶ Pow B of a predicate `φ : B ⟶ Ω C`. 150 | This is the global element of `Pow B` associated to a predicate 151 | on `B`. 152 | -/ 153 | def name {B} (φ : B ⟶ Ω C) : ⊤_ C ⟶ pow B := (((prod.fst) ≫ φ))^ 154 | 155 | /-- Notation for `Name`. Consistent with Mac Lane and Moerdijk's notation. -/ 156 | notation "⌈" φ "⌉" => name φ 157 | 158 | /-- The inverse of `Name`, sending a global element of `Pow B` 159 | to the corresponding predicate on `B`. 160 | -/ 161 | def Predicate.fromName {B} (φ' : ⊤_ C ⟶ pow B) : B ⟶ Ω C := 162 | (prod.lift (𝟙 B) (terminal.from B)) ≫ φ'^ 163 | 164 | /-- The condition from the definition of `Name` 165 | as the `P_transpose` of a morphism. 166 | -/ 167 | lemma Predicate.NameDef {B} (φ : B ⟶ Ω C) : (prod.map (𝟙 _) ⌈φ⌉) ≫ (in_ B) = (prod.fst) ≫ φ := 168 | comm _ _ 169 | 170 | /-- The equivalence between morphisms `B ⟶ Ω C` and morphisms `⊤_ C ⟶ pow B`, 171 | which comes from the more general equivalence between morphisms `B ⨯ A ⟶ Ω C` 172 | and morphisms `A ⟶ pow B`. 173 | -/ 174 | def Predicate.NameEquiv (B : C) : (B ⟶ Ω C) ≃ (⊤_ C ⟶ pow B) where 175 | toFun := name 176 | invFun := fromName 177 | left_inv := by 178 | intro φ 179 | dsimp [name, fromName] 180 | rw [transpose_left_inv, ←assoc, prod.lift_fst, id_comp] 181 | right_inv := by 182 | intro φ' 183 | dsimp only [name, fromName] 184 | have h := (Limits.prod.rightUnitor B).hom_inv_id 185 | dsimp at h 186 | rw [←assoc, h, id_comp, transpose_right_inv] 187 | 188 | end 189 | end CategoryTheory.Topos 190 | -------------------------------------------------------------------------------- /Topos/Power.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Charlie Conneen. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Charlie Conneen 5 | -/ 6 | --import Mathlib.CategoryTheory.Topos.Classifier 7 | import Mathlib.CategoryTheory.ChosenFiniteProducts 8 | import Topos.Classifier 9 | 10 | /-! 11 | # Power Objects 12 | 13 | Defines power objects for a category C with a subobject classifier and pullbacks. 14 | 15 | ## Main definitions 16 | 17 | Let `C` be a category with a terminal object, a subobject classifier, and pullbacks. 18 | 19 | * For an object `B : C`, `CategoryTheory.Power.PowerObject B` contains the data 20 | of a power object for `B`. 21 | 22 | * `CategoryTheory.Power.HasPowerObjects C` says that every object in `C` has a power object. 23 | 24 | * If `C` has all power objects, then there is a functor `powFunctor C : Cᵒᵖ ⥤ C` which 25 | sends an object `B : C` to its power object `pow B`, and sends a morphism `f : A ⟶ B` to the 26 | corresponding "inverse image" morphism `inverseImage f : pow B ⟶ pow A`. 27 | 28 | ## Main results 29 | 30 | * `powFunctor C` is self adjoint, in the sense that its adjoint is the corresponding 31 | functor `powFunctorOp C : C ⥤ Cᵒᵖ`. This is exhibited as `powSelfAdj C`. 32 | 33 | ## Notation 34 | 35 | * if `f : B ⨯ A ⟶ Ω C` is a morphism in a category with power objects, then 36 | `f^` denotes the corresponding morphism `A ⟶ pow B`. 37 | * If `g : A ⟶ pow B` is a morphism in a category with power objects, then 38 | `g^` denotes the corresponding morphism `B ⨯ A ⟶ Ω C`. 39 | * To "curry" maps in the other coordinate, put the caret `^` before the function argument 40 | instead of after. 41 | 42 | ## References 43 | 44 | * [S. MacLane and I. Moerdijk, *Sheaves in Geometry and Logic*][MLM92] 45 | 46 | -/ 47 | 48 | 49 | open CategoryTheory Category Limits HasClassifier MonoidalCategory 50 | 51 | namespace CategoryTheory.Power 52 | 53 | universe u v 54 | variable {C : Type u} [Category.{v} C] [HasTerminal C] [HasClassifier C] 55 | variable [ChosenFiniteProducts C] 56 | 57 | variable (A B : C) 58 | 59 | #synth MonoidalCategory C 60 | 61 | #check A ⊗ B 62 | 63 | /-- An object `PB` and a map `in_B : B ⨯ PB ⟶ Ω C` form a power object for `B : C` 64 | if, for any map `f : B ⨯ A ⟶ Ω C`, there is a unique map `f' : A ⟶ PB` such that 65 | the following diagram commutes: 66 | ``` 67 | 68 | B ⨯ A ---f---> Ω C 69 | | ^ 70 | | / 71 | | / 72 | (𝟙 B) ⨯ f' / 73 | | / 74 | | in_B 75 | | / 76 | | / 77 | | / 78 | | / 79 | v / 80 | B ⨯ PB 81 | ``` 82 | -/ 83 | structure PowerObject (B : C) where 84 | /-- The power object -/ 85 | pow : C 86 | /-- The membership predicate -/ 87 | in_ : B ⨯ pow ⟶ Ω C 88 | /-- The transpose of a map -/ 89 | transpose {A} (f : B ⨯ A ⟶ Ω C) : A ⟶ pow 90 | /-- the characterizing property of the transpose -/ 91 | comm {A} (f : B ⨯ A ⟶ Ω C) : prod.map (𝟙 B) (transpose f) ≫ in_ = f 92 | /-- `transpose f` is the only map which satisfies the commutativity condition-/ 93 | uniq {A} {f : B ⨯ A ⟶ Ω C} {hat' : A ⟶ pow} (hat'_comm : prod.map (𝟙 B) hat' ≫ in_ = f) : transpose f = hat' 94 | 95 | variable (C) 96 | 97 | /-- A category has power objects if each of its objects has a power object. -/ 98 | class HasPowerObjects : Prop where 99 | /-- Each `B : C` has a power object. -/ 100 | has_power_object (B : C) : Nonempty (PowerObject B) 101 | 102 | variable {C} 103 | 104 | variable [HasPowerObjects C] 105 | 106 | noncomputable section 107 | 108 | /-- Notation for the power object of an object. -/ 109 | abbrev pow (B : C) : C := (HasPowerObjects.has_power_object B).some.pow 110 | 111 | /-- Notation for the predicate "b ∈ S" as a map `B ⨯ pow B ⟶ Ω`. -/ 112 | abbrev in_ (B : C) : B ⨯ pow B ⟶ Ω C := (HasPowerObjects.has_power_object B).some.in_ 113 | 114 | /-- The map Hom(B⨯A,Ω) → Hom(A,P(B)). 115 | This is currying the left argument. 116 | -/ 117 | def transpose {B A} (f : B ⨯ A ⟶ Ω C) : A ⟶ pow B := 118 | (HasPowerObjects.has_power_object B).some.transpose f 119 | 120 | /-- Shorthand for currying the left argument. -/ 121 | notation f "^" => transpose f 122 | 123 | /-- The characterizing property of `transpose`. -/ 124 | @[reassoc] 125 | lemma comm (B) {A} (f : B ⨯ A ⟶ Ω C) : prod.map (𝟙 _) (f^) ≫ in_ B = f := 126 | (HasPowerObjects.has_power_object B).some.comm f 127 | 128 | /-- `transpose` is the only map which satisfies the commutativity condition. -/ 129 | lemma unique (B) {A} {f : B ⨯ A ⟶ Ω C} {hat' : A ⟶ pow B} (hat'_comm : prod.map (𝟙 _) hat' ≫ in_ B = f ) : 130 | f^ = hat' := 131 | (HasPowerObjects.has_power_object B).some.uniq hat'_comm 132 | 133 | /-- Un-currying on the left. -/ 134 | abbrev transposeInv {B A} (f : A ⟶ pow B) : B ⨯ A ⟶ Ω C := 135 | (prod.map (𝟙 _) f) ≫ in_ B 136 | 137 | /-- Shorthand for un-currying on the left. -/ 138 | notation f "^" => transposeInv f 139 | 140 | /-- Equivalence between Hom(B⨯A,Ω) and Hom(A,P(B)). -/ 141 | def transposeEquiv (A B : C) : (B ⨯ A ⟶ Ω C) ≃ (A ⟶ pow B) where 142 | toFun := transpose 143 | invFun := transposeInv 144 | left_inv := fun f => comm _ _ 145 | right_inv := by 146 | intro g 147 | apply unique 148 | rfl 149 | 150 | /-- `transposeInv` is a left inverse of `transpose`. -/ 151 | @[reassoc (attr := simp)] 152 | lemma transpose_left_inv {B A} (f : B ⨯ A ⟶ Ω C) : (f^)^ = f := 153 | (transposeEquiv _ _).left_inv _ 154 | 155 | /-- `transposeInv` is a right inverse of `transpose`. -/ 156 | @[reassoc (attr := simp)] 157 | lemma transpose_right_inv {B A : C} (f : A ⟶ pow B) : (f^)^ = f := 158 | (transposeEquiv _ _).right_inv _ 159 | 160 | /-- The map Hom(B⨯A,Ω) → Hom(B,P(A)). 161 | This is currying the right argument. 162 | -/ 163 | def transposeSymm {B A} (f : B ⨯ A ⟶ Ω C) : B ⟶ pow A := 164 | transpose ((prod.braiding A B).hom ≫ f) 165 | 166 | /-- Shorthand for currying the right argument. -/ 167 | notation "^" f => transposeSymm f 168 | 169 | /-- Un-currying on the right. -/ 170 | abbrev transposeSymmInv {B A} (f : B ⟶ pow A) : B ⨯ A ⟶ Ω C := 171 | (prod.braiding A B).inv ≫ (transposeInv f) 172 | 173 | /-- Shorthand for un-currying on the right. -/ 174 | notation "^" f => transposeSymmInv f 175 | 176 | /-- Equivalence between Hom(B⨯A,Ω) and Hom(B,P(A)). -/ 177 | def transposeEquivSymm (A B : C) : (B ⨯ A ⟶ Ω C) ≃ (B ⟶ pow A) where 178 | toFun := transposeSymm 179 | invFun := transposeSymmInv 180 | left_inv := by 181 | intro f 182 | dsimp only [transposeSymm, transposeInv, transposeSymmInv] 183 | rw [comm, ←assoc, Iso.inv_hom_id, id_comp] 184 | right_inv := by 185 | intro g 186 | apply unique 187 | rw [←assoc, Iso.hom_inv_id, id_comp] 188 | 189 | /-- `transposeSymmInv` is the left inverse 190 | of `transposeSymm`. 191 | -/ 192 | @[simp] 193 | lemma transpose_symm_left_inv {B A} (f : B ⨯ A ⟶ Ω C) : 194 | (^(^f)) = f := 195 | (transposeEquivSymm _ _).left_inv _ 196 | 197 | /-- `transposeSymmInv` is the right inverse 198 | of `transposeSymm`. 199 | -/ 200 | @[simp] 201 | lemma transpose_symm_right_inv {B A : C} (f : B ⟶ pow A) : 202 | (^(^f)) = f := 203 | (transposeEquivSymm _ _).right_inv _ 204 | 205 | /-- Equivalence between Hom(A,P(B)) and Hom(B, P(A)). 206 | This is just the composition of `transposeEquiv` and `transposeEquivSymm`. 207 | -/ 208 | def transpose_transpose_Equiv (A B : C) : (B ⟶ pow A) ≃ (A ⟶ pow B) := 209 | -- (transposeEquivSymm A B).symm.trans (transposeEquiv A B) 210 | Equiv.trans (transposeEquivSymm A B).symm (transposeEquiv A B) 211 | 212 | 213 | /-- 214 | The power object functor's action on arrows. 215 | Sends `h : A ⟶ B` to the P-transpose of the map `h⨯1 ≫ ∈_B : A ⨯ pow B ⟶ B ⨯ pow B ⟶ Ω`. 216 | -/ 217 | def inverseImage {B A : C} (h : A ⟶ B) : pow B ⟶ pow A := 218 | ((prod.map h (𝟙 (pow B))) ≫ (in_ B))^ 219 | 220 | /-- The following diagram commutes: 221 | ``` 222 | A ⨯ Pow B ----(𝟙 A) ⨯ Pow_map h----> A ⨯ Pow A 223 | | | 224 | | | 225 | h ⨯ (𝟙 (Pow B)) in_ A 226 | | | 227 | v v 228 | B ⨯ Pow B ----------in_ B-----------> Ω C 229 | ``` 230 | -/ 231 | lemma inverseImage_comm {A B : C} (h : A ⟶ B) : 232 | (prod.map (𝟙 A) (inverseImage h)) ≫ in_ A = (prod.map h (𝟙 (pow B))) ≫ (in_ B) := by 233 | dsimp [inverseImage] 234 | apply comm 235 | 236 | /-- `powMap` sends the identity on an object `X` to the identity on `pow X`. -/ 237 | @[simp] 238 | lemma inverseImage_id {B : C} : inverseImage (𝟙 B) = 𝟙 (pow B) := by 239 | apply unique; rfl 240 | 241 | 242 | variable (C) 243 | 244 | /-- 245 | The power object functor. 246 | Sends objects `B` to their power objects `pow B`. 247 | Sends arrows `h : A ⟶ B` to the P-transpose of the map `h⨯1 ≫ ∈_B : A ⨯ pow B ⟶ B ⨯ pow B ⟶ Ω`, 248 | which is the "preimage" morphism `pow B ⟶ pow A`. 249 | -/ 250 | def powFunctor : Cᵒᵖ ⥤ C where 251 | obj := fun ⟨B⟩ ↦ pow B 252 | map := fun ⟨h⟩ ↦ inverseImage h 253 | map_id := by 254 | intro _ 255 | apply unique 256 | rfl 257 | map_comp := by 258 | intro ⟨X⟩ ⟨Y⟩ ⟨Z⟩ ⟨f⟩ ⟨g⟩ 259 | apply unique 260 | dsimp [inverseImage] 261 | symm 262 | calc 263 | prod.map (g ≫ f) (𝟙 (pow X)) ≫ in_ X 264 | = (prod.map g (𝟙 (pow X))) ≫ (prod.map f (𝟙 (pow X))) ≫ in_ X := by 265 | rw [←assoc, ←prod.map_comp_id] 266 | _ = (prod.map g (𝟙 (pow X))) ≫ (prod.map (𝟙 Y) (inverseImage f)) ≫ in_ Y := by 267 | rw [inverseImage_comm] 268 | _ = (prod.map (𝟙 Z) (inverseImage f)) ≫ (prod.map g (𝟙 (pow Y))) ≫ in_ Y := by 269 | repeat rw [prod.map_map_assoc, comp_id, id_comp] 270 | _ = (prod.map (𝟙 Z) (inverseImage f)) ≫ (prod.map (𝟙 Z) (inverseImage g)) ≫ in_ Z := by 271 | rw [inverseImage_comm] 272 | _ = prod.map (𝟙 Z) (inverseImage f ≫ inverseImage g ) ≫ in_ Z := by 273 | rw [←assoc, prod.map_id_comp] 274 | 275 | /-- The power object functor, treated as a functor `C ⥤ Cᵒᵖ`. -/ 276 | def powFunctorOp : C ⥤ Cᵒᵖ where 277 | obj := fun B ↦ ⟨pow B⟩ 278 | map := fun h ↦ ⟨inverseImage h⟩ 279 | map_id := by 280 | intro _ 281 | apply congrArg Opposite.op 282 | apply (powFunctor C).map_id 283 | map_comp := by 284 | intros 285 | apply congrArg Opposite.op 286 | apply (powFunctor C).map_comp 287 | 288 | /-- exhibiting that the pow functor is adjoint to itself on the right. -/ 289 | def powSelfAdj : powFunctorOp C ⊣ powFunctor C := by 290 | apply Adjunction.mkOfHomEquiv 291 | fapply Adjunction.CoreHomEquiv.mk 292 | 293 | -- homEquiv step 294 | · exact fun X ⟨Y⟩ => { 295 | toFun := fun ⟨f⟩ => (transpose_transpose_Equiv X Y).toFun f 296 | invFun := fun g => ⟨(transpose_transpose_Equiv X Y).invFun g⟩ 297 | left_inv := fun ⟨f⟩ => by simp 298 | right_inv := fun g => by simp 299 | } 300 | 301 | -- homEquiv_naturality_left_symm step 302 | · intro X' X ⟨Y⟩ f g 303 | simp 304 | congr 305 | show (transpose_transpose_Equiv X' Y).symm (f ≫ g) = 306 | (transpose_transpose_Equiv X Y).symm g ≫ inverseImage f 307 | dsimp only [transpose_transpose_Equiv, transposeEquivSymm, transposeEquiv] 308 | simp 309 | dsimp only [transposeSymm, transposeInv, inverseImage] 310 | apply unique 311 | rw [prod.map_id_comp _ (transpose _), assoc _ _ (in_ X'), comm, ←assoc _ _ (in_ X), 312 | prod.map_map, id_comp, comp_id, ←comp_id f, ←id_comp (transpose _), ←prod.map_map, 313 | assoc, comm] 314 | have h : prod.map f (𝟙 Y) ≫ (prod.braiding X Y).hom = (prod.braiding _ _).hom ≫ prod.map (𝟙 _) f := by simp 315 | rw [←assoc (prod.map f (𝟙 _)), h] 316 | simp 317 | 318 | -- homEquiv_naturality_right step 319 | · intro X ⟨Y⟩ ⟨Y'⟩ ⟨f⟩ ⟨g⟩ 320 | dsimp only [transpose_transpose_Equiv, transposeEquiv, transposeEquivSymm] 321 | simp only [prod.lift_map_assoc, comp_id, Equiv.toFun_as_coe, Equiv.trans_apply, 322 | Equiv.coe_fn_symm_mk, Equiv.coe_fn_mk, Equiv.invFun_as_coe, Equiv.symm_trans_apply, 323 | Equiv.symm_symm] 324 | show transpose ((prod.braiding X Y').inv ≫ prod.map (𝟙 X) (g ≫ f) ≫ in_ X) = 325 | transpose ((prod.braiding X Y).inv ≫ prod.map (𝟙 X) f ≫ in_ X) ≫ inverseImage g 326 | dsimp only [transposeInv, inverseImage] 327 | apply unique 328 | rw [prod.map_id_comp (transpose _), assoc, comm, ←assoc _ _ (in_ Y), 329 | prod.map_map, id_comp, comp_id, ←comp_id g] 330 | have h : prod.map g (𝟙 X) ≫ (prod.braiding X Y).inv = (prod.braiding _ _).inv ≫ prod.map (𝟙 _) g := by simp 331 | rw [←id_comp (transpose _), ←prod.map_map, assoc, comm, ←assoc (prod.map g _), h] 332 | simp only [prod.braiding_inv, prod.lift_map_assoc, comp_id, prod.lift_map, assoc] 333 | 334 | 335 | end 336 | end CategoryTheory.Power 337 | -------------------------------------------------------------------------------- /Topos/Classifier.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Charlie Conneen. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Charlie Conneen 5 | -/ 6 | import Mathlib.CategoryTheory.Limits.Shapes.Pullback.CommSq 7 | import Mathlib.CategoryTheory.Limits.Shapes.RegularMono 8 | import Mathlib.Tactic.ApplyFun 9 | import Mathlib.CategoryTheory.Subobject.Basic 10 | 11 | /-! 12 | 13 | # Subobject Classifier 14 | 15 | We define what it means for a morphism in a category to be a subobject 16 | classifier as `CategoryTheory.HasClassifier`. The reason for this 17 | naming convention is to distinguish this internal categorical notion of 18 | a subobject classifier from a classifier of terms of type `Subobject B` 19 | for `B : C`. 20 | 21 | c.f. the following Lean 3 code, where similar work was done: 22 | https://github.com/b-mehta/topos/blob/master/src/subobject_classifier.lean 23 | 24 | ## Main definitions 25 | 26 | Let `C` refer to a category with a terminal object. 27 | 28 | * `CategoryTheory.Classifier C` is the data of a subobject classifier 29 | in `C`. 30 | 31 | * `CategoryTheory.HasClassifier C` says that there is at least one 32 | subobject classifier. 33 | 34 | ## Main results 35 | 36 | * It is a theorem that the truth morphism `⊤_ C ⟶ Ω C` is a (split, and 37 | therefore regular) monomorphism. 38 | 39 | * `MonoClassifier.balanced` shows that any category with a subobject classifier 40 | is balanced. This follows from the fact that every monomorphism is the 41 | pullback of a regular monomorphism (the truth morphism). 42 | 43 | ## Notation 44 | 45 | * if `m` is a monomorphism, `χ_ m` denotes characteristic map of `m`, 46 | which is the corresponding map to the subobject classifier. 47 | 48 | ## References 49 | 50 | * [S. MacLane and I. Moerdijk, *Sheaves in Geometry and Logic*][MLM92] 51 | 52 | -/ 53 | 54 | universe u v u₀ v₀ 55 | 56 | open CategoryTheory Category Limits Functor 57 | 58 | variable (C : Type u) [Category.{v} C] [HasTerminal C] 59 | 60 | namespace CategoryTheory 61 | 62 | /-- A morphism `t : ⊤_ C ⟶ Ω` from the terminal object of a category `C` 63 | is a subobject classifier if, for every monomorphism `m : U ⟶ X` in `C`, 64 | there is a unique map `χ : X ⟶ Ω` such that the following square is a pullback square: 65 | ``` 66 | U ---------m----------> X 67 | | | 68 | terminal.from U χ 69 | | | 70 | v v 71 | ⊤_ C --------t----------> Ω 72 | ``` 73 | -/ 74 | structure Classifier where 75 | /-- The target of the truth morphism -/ 76 | {Ω : C} 77 | /-- the truth morphism for a subobject classifier -/ 78 | t : ⊤_ C ⟶ Ω 79 | /-- For any monomorphism `U ⟶ X`, there is an associated characteristic map `X ⟶ Ω`. -/ 80 | char {U X : C} (m : U ⟶ X) [Mono m] : X ⟶ Ω 81 | /-- `char m` forms the appropriate pullback square. -/ 82 | isPullback {U X : C} (m : U ⟶ X) [Mono m] : IsPullback m (terminal.from U) (char m) t 83 | /-- `char m` is the only map `X ⟶ Ω` which forms the appropriate pullback square. -/ 84 | uniq {U X : C} (m : U ⟶ X) [Mono m] (χ : X ⟶ Ω) (hχ : IsPullback m (terminal.from U) χ t) : 85 | χ = char m 86 | 87 | 88 | /-- A category `C` has a subobject classifier if there is at least one subobject classifier. -/ 89 | class HasClassifier : Prop where 90 | /-- There is some classifier. -/ 91 | exists_classifier : Nonempty (Classifier C) 92 | 93 | namespace HasClassifier 94 | 95 | variable [HasClassifier C] 96 | 97 | noncomputable section 98 | 99 | /-- Notation for the object in an arbitrary choice of a subobject classifier -/ 100 | abbrev Ω : C := HasClassifier.exists_classifier.some.Ω 101 | 102 | /-- Notation for the "truth arrow" in an arbitrary choice of a subobject classifier -/ 103 | abbrev t : ⊤_ C ⟶ Ω C := HasClassifier.exists_classifier.some.t 104 | 105 | variable {C} 106 | variable {U X : C} (m : U ⟶ X) [Mono m] 107 | 108 | /-- returns the characteristic morphism of the subobject `(m : U ⟶ X) [Mono m]` -/ 109 | def characteristicMap : X ⟶ Ω C := 110 | HasClassifier.exists_classifier.some.char m 111 | 112 | /-- shorthand for the characteristic morphism, `ClassifierOf m` -/ 113 | abbrev χ_ := characteristicMap m 114 | 115 | /-- The diagram 116 | ``` 117 | U ---------m----------> X 118 | | | 119 | terminal.from U χ_ m 120 | | | 121 | v v 122 | ⊤_ C -------t C---------> Ω 123 | ``` 124 | is a pullback square. 125 | -/ 126 | lemma isPullback : IsPullback m (terminal.from U) (χ_ m) (t C) := 127 | HasClassifier.exists_classifier.some.isPullback m 128 | 129 | /-- The diagram 130 | ``` 131 | U ---------m----------> X 132 | | | 133 | terminal.from U χ_ m 134 | | | 135 | v v 136 | ⊤_ C -------t C---------> Ω 137 | ``` 138 | commutes. 139 | -/ 140 | @[reassoc] 141 | lemma comm : m ≫ (χ_ m) = terminal.from _ ≫ t C := (isPullback m).w 142 | 143 | /-- `characteristicMap m` is the only map for which the associated square 144 | is a pullback square. 145 | -/ 146 | lemma unique (χ : X ⟶ Ω C) (hχ : IsPullback m (terminal.from _) χ (t C)) : χ = χ_ m := 147 | HasClassifier.exists_classifier.some.uniq m χ hχ 148 | 149 | /-- `t C` is a regular monomorphism (because it is split). -/ 150 | noncomputable instance truthIsRegularMono : RegularMono (t C) := 151 | RegularMono.ofIsSplitMono (t C) 152 | 153 | /-- The following diagram 154 | ``` 155 | U ---------m----------> X 156 | | | 157 | terminal.from U χ_ m 158 | | | 159 | v v 160 | ⊤_ C -------t C---------> Ω 161 | ``` 162 | being a pullback for any monic `m` means that every monomorphism 163 | in `C` is the pullback of a regular monomorphism; since regularity 164 | is stable under base change, every monomorphism is regular. 165 | -/ 166 | noncomputable instance monoIsRegularMono {A B : C} (m : A ⟶ B) [Mono m] : RegularMono m := 167 | regularOfIsPullbackFstOfRegular (isPullback m).w (isPullback m).isLimit 168 | 169 | /-- `C` is a balanced category. -/ 170 | instance balanced : Balanced C where 171 | isIso_of_mono_of_epi := fun f => isIso_of_epi_of_strongMono f 172 | 173 | instance : Balanced Cᵒᵖ := balanced_opposite 174 | 175 | /-- If the source of a faithful functor has a subobject classifier, the functor reflects 176 | isomorphisms. This holds for any balanced category. 177 | -/ 178 | instance reflectsIsomorphisms (D : Type u₀) [Category.{v₀} D] (F : C ⥤ D) [Functor.Faithful F] : 179 | Functor.ReflectsIsomorphisms F := 180 | reflectsIsomorphisms_of_reflectsMonomorphisms_of_reflectsEpimorphisms F 181 | 182 | /-- If the source of a faithful functor is the opposite category of one with a subobject classifier, 183 | the same holds -- the functor reflects isomorphisms. 184 | -/ 185 | instance reflectsIsomorphismsOp (D : Type u₀) [Category.{v₀} D] 186 | (F : Cᵒᵖ ⥤ D) [Functor.Faithful F] : 187 | Functor.ReflectsIsomorphisms F := 188 | reflectsIsomorphisms_of_reflectsMonomorphisms_of_reflectsEpimorphisms F 189 | 190 | variable [HasPullbacks C] 191 | variable (classifier : Classifier C) 192 | 193 | variable (B : C) (φ : B ⟶ classifier.Ω) 194 | variable (s : Cone (cospan φ classifier.t)) 195 | 196 | example : (Subobject.mk (pullback.fst φ classifier.t)).arrow = 197 | (Subobject.underlyingIso (pullback.fst φ classifier.t)).hom ≫ pullback.fst φ classifier.t := by 198 | exact Eq.symm (Subobject.underlyingIso_hom_comp_eq_mk (pullback.fst φ classifier.t)) 199 | 200 | example : pullback.fst φ classifier.t ≫ φ = pullback.snd φ classifier.t ≫ classifier.t := by 201 | exact pullback.condition 202 | 203 | -- TODO: refactor so this `omit` is unnecessary. 204 | omit [HasClassifier C] in 205 | lemma subobject_pullback {B : C} (φ : B ⟶ classifier.Ω) : 206 | IsPullback (Subobject.mk (pullback.fst φ classifier.t)).arrow 207 | (terminal.from (Subobject.underlying.obj (Subobject.mk (pullback.fst φ classifier.t)))) 208 | φ classifier.t := by 209 | apply IsPullback.of_iso_pullback ?_ (Subobject.underlyingIso (pullback.fst φ classifier.t)) 210 | · exact Subobject.underlyingIso_hom_comp_eq_mk (pullback.fst φ classifier.t) 211 | · exact terminal.hom_ext 212 | ((Subobject.underlyingIso (pullback.fst φ classifier.t)).hom ≫ pullback.snd φ classifier.t) 213 | (terminal.from (Subobject.underlying.obj (Subobject.mk (pullback.fst φ classifier.t)))) 214 | · apply CommSq.mk 215 | rw [(Subobject.underlyingIso_hom_comp_eq_mk (pullback.fst φ classifier.t)).symm, 216 | assoc, pullback.condition, ←assoc] 217 | have h : ((Subobject.underlyingIso (pullback.fst φ classifier.t)).hom 218 | ≫ pullback.snd φ classifier.t) = terminal.from _ := by apply terminal.hom_ext 219 | rw [h] 220 | 221 | variable (S : Subobject B) 222 | 223 | def subobjectIso1 : Subobject.underlying.obj 224 | (Subobject.mk (pullback.fst (classifier.char S.arrow) classifier.t)) 225 | ≅ pullback (classifier.char S.arrow) classifier.t := 226 | Subobject.underlyingIso (pullback.fst (classifier.char S.arrow) classifier.t) 227 | 228 | def subobjectIso2 : 229 | pullback (classifier.char S.arrow) classifier.t ≅ Subobject.underlying.obj S where 230 | hom := 231 | (classifier.isPullback S.arrow).lift 232 | (pullback.fst (classifier.char S.arrow) classifier.t) 233 | (pullback.snd (classifier.char S.arrow) classifier.t) 234 | pullback.condition 235 | inv := pullback.lift S.arrow (terminal.from _) (classifier.isPullback S.arrow).w 236 | 237 | def subobjectIso3 : 238 | Subobject.underlying.obj (Subobject.mk (pullback.fst (classifier.char S.arrow) classifier.t)) ≅ 239 | Subobject.underlying.obj S where 240 | hom := (subobjectIso1 classifier B S).hom ≫ (subobjectIso2 classifier B S).hom 241 | inv := (subobjectIso2 classifier B S).inv ≫ (subobjectIso1 classifier B S).inv 242 | 243 | def subobjectEquiv (B : C) : (B ⟶ classifier.Ω) ≃ Subobject B where 244 | toFun := fun φ => Subobject.mk (pullback.fst φ classifier.t) 245 | invFun := fun S => classifier.char S.arrow 246 | left_inv := by 247 | intro φ 248 | simp 249 | symm 250 | apply classifier.uniq 251 | exact subobject_pullback classifier φ 252 | right_inv := by 253 | intro S 254 | ext 255 | · exact subobjectIso3 classifier B S 256 | · simp [subobjectIso3, subobjectIso1, subobjectIso2] 257 | 258 | def classifierHomEquiv (B : C) : (B ⟶ classifier.Ω) ≃ (B ⟶ Ω C) where 259 | toFun := fun φ ↦ 260 | (subobjectEquiv HasClassifier.exists_classifier.some _).invFun ( 261 | (subobjectEquiv classifier _).toFun φ) 262 | invFun := fun φ' ↦ 263 | (subobjectEquiv classifier _).invFun ( 264 | (subobjectEquiv HasClassifier.exists_classifier.some _).toFun φ') 265 | left_inv := by intro φ; simp 266 | right_inv := by intro φ'; simp 267 | 268 | variable {Z Z' : C} (f : Z' ⟶ Z) (g : Z ⟶ classifier.Ω) 269 | 270 | instance : Mono (classifier.t) := IsSplitMono.mono classifier.t 271 | 272 | def classifierIso (classifier : Classifier C) : classifier.Ω ≅ Ω C := by 273 | apply Yoneda.ext ( 274 | p := fun φ ↦ (classifierHomEquiv classifier _).toFun φ) ( 275 | q := fun φ' ↦ (classifierHomEquiv classifier _).invFun φ') 276 | · simp 277 | · simp 278 | · intro Z Z' f g 279 | dsimp 280 | /- 281 | show (subobjectEquiv HasClassifier.exists_classifier.some _).invFun ( 282 | (subobjectEquiv classifier _).toFun (f ≫ g)) 283 | = f ≫ (subobjectEquiv HasClassifier.exists_classifier.some _).invFun ( 284 | (subobjectEquiv classifier _).toFun g) 285 | -/ 286 | show χ_ (Subobject.mk (pullback.fst (f ≫ g) classifier.t)).arrow 287 | = f ≫ χ_ (Subobject.mk (pullback.fst g (classifier.t))).arrow 288 | symm 289 | apply HasClassifier.unique 290 | apply IsPullback.of_iso_pullback 291 | · apply CommSq.mk 292 | /- 293 | 294 | -/ 295 | let l' : Subobject.underlying.obj (Subobject.mk (pullback.fst (f ≫ g) classifier.t)) 296 | ≅ pullback (f ≫ g) classifier.t := 297 | Subobject.underlyingIso (pullback.fst (f ≫ g) classifier.t) 298 | 299 | have h : (Subobject.mk (pullback.fst (f ≫ g) classifier.t)).arrow 300 | = (Subobject.underlyingIso (pullback.fst (f ≫ g) classifier.t)).hom 301 | ≫ (pullback.fst (f ≫ g) classifier.t) := 302 | Eq.symm (Subobject.underlyingIso_hom_comp_eq_mk (pullback.fst (f ≫ g) classifier.t)) 303 | rw [(Subobject.underlyingIso_hom_comp_eq_mk (pullback.fst (f ≫ g) classifier.t)).symm] 304 | rw [assoc] 305 | let lift : pullback (f ≫ g) classifier.t ⟶ pullback g classifier.t := by 306 | refine pullback.lift (pullback.fst (f ≫ g) classifier.t ≫ f) (terminal.from _) ?_ 307 | have h' : terminal.from (pullback (f ≫ g) classifier.t) = pullback.snd _ _ := by 308 | apply terminal.hom_ext 309 | rw [assoc, pullback.condition, h'] 310 | have h' : (pullback.fst _ _) ≫ f = lift ≫ (pullback.fst _ _) := by 311 | symm; apply pullback.lift_fst 312 | rw [←(assoc _ f), h', assoc] 313 | -- let lift₁ : 314 | sorry 315 | · sorry 316 | · sorry 317 | · sorry 318 | 319 | 320 | end 321 | end CategoryTheory.HasClassifier 322 | -------------------------------------------------------------------------------- /Topos/Exponentials.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Charlie Conneen. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Charlie Conneen 5 | -/ 6 | import Mathlib.CategoryTheory.Monoidal.OfHasFiniteProducts 7 | --import Mathlib.CategoryTheory.Topos.Basic 8 | import Mathlib.CategoryTheory.Closed.Cartesian 9 | import Topos.Basic 10 | 11 | /-! 12 | # Exponential Objects 13 | 14 | Proves that a topos has exponential objects (internal homs). 15 | Consequently, every topos is Cartesian closed. 16 | 17 | ## Main definitions 18 | 19 | * `Hom A B` is the exponential object, and `eval A B` is the associated 20 | "evaluation map" `A ⨯ Hom A B ⟶ B`. 21 | 22 | * `IsExponentialObject` says what it means to be an exponential object. 23 | 24 | ## Main results 25 | 26 | * `ToposHasExponentials` shows that a topos has exponential objects. 27 | This is done by showing `IsExponentialObject (eval A B)`. 28 | 29 | * `ExpAdjEquiv` exhibits `(A ⨯ X ⟶ B) ≃ (X ⟶ Hom A B)` for any `A B X : C` 30 | in a topos `C`. 31 | 32 | ## References 33 | 34 | * [S. MacLane and I. Moerdijk, *Sheaves in Geometry and Logic*][MLM92] 35 | 36 | -/ 37 | 38 | open CategoryTheory Category Limits MonoClassifier Power Topos 39 | 40 | universe u v 41 | 42 | variable {C : Type u} [Category.{v} C] 43 | 44 | namespace CategoryTheory.Topos 45 | 46 | noncomputable section 47 | 48 | variable [IsTopos C] 49 | 50 | /-- The exponential object B^A. -/ 51 | def hom (A B : C) : C := 52 | pullback 53 | ((((prod.associator _ _ _).inv ≫ in_ (B ⨯ A))^ ≫ Predicate.isSingleton B)^) 54 | ⌈Predicate.true_ A⌉ 55 | 56 | /-- The map which sends `A ⟶ B` to its graph as a subobject of `B ⨯ A`. -/ 57 | def homToGraph (A B : C) : hom A B ⟶ pow (B ⨯ A) := pullback.fst _ _ 58 | 59 | /-- The map sending a morphism to its graph is a monomorphism, so that 60 | the graphs of morphisms `A ⟶ B` may be regarded as subobjects of `B ⨯ A`. 61 | -/ 62 | instance homToGraphMono {A B : C} : Mono (homToGraph A B) := pullback.fst_of_mono 63 | 64 | /-- Convenience lemma used in `Hom_comm`. -/ 65 | lemma exp_cone_snd (A B : C) : 66 | pullback.snd _ _ = terminal.from (hom A B) := Unique.eq_default _ 67 | 68 | /-- Convenience lemma used in `EvalDef_comm`. -/ 69 | lemma hom_comm (A B : C) : 70 | homToGraph A B ≫ ((((prod.associator _ _ _).inv ≫ in_ (B ⨯ A))^ ≫ Predicate.isSingleton B)^) 71 | = terminal.from (hom A B) ≫ ⌈Predicate.true_ A⌉ := by 72 | rw [←exp_cone_snd]; exact pullback.condition 73 | 74 | /-- This lemma states that the following diagram commutes: 75 | ``` 76 | A ⨯ Hom A B ---------terminal.from _----------> ⊤_ C 77 | | | 78 | | | 79 | (𝟙 A) ⨯ homToGraph A B t C 80 | | | 81 | | | 82 | v v 83 | A ⨯ Pow (B ⨯ A) ------------u-------------------> Ω 84 | ``` 85 | where `u` intuitively is the predicate: 86 | (a,S) ↦ "there is exactly one b in B such that (b,a) in S". 87 | This is used to define the map `eval A B : A ⨯ Hom A B ⟶ B` 88 | as a `pullback.lift` where the object `B` serves as the pullback. 89 | -/ 90 | lemma eval_def_comm (A B : C) : 91 | (prod.map (𝟙 A) (homToGraph A B) 92 | ≫ ((prod.associator _ _ _).inv ≫ in_ (B ⨯ A))^) ≫ Predicate.isSingleton B 93 | = Predicate.true_ (A ⨯ hom A B) := by 94 | let id_m : A ⨯ hom A B ⟶ A ⨯ pow (B ⨯ A) := prod.map (𝟙 _) (homToGraph A B) 95 | let v := (((prod.associator _ _ _).inv ≫ in_ (B ⨯ A))^) 96 | let σ_B := Predicate.isSingleton B 97 | let u := ((v ≫ Predicate.isSingleton B)^) 98 | let id_u := prod.map (𝟙 A) u 99 | have comm_middle : v ≫ σ_B = id_u ≫ (in_ A) := (comm A (v ≫ σ_B)).symm 100 | have comm_left : 101 | id_m ≫ id_u 102 | = prod.map (𝟙 _) (terminal.from _) ≫ prod.map (𝟙 _) ⌈Predicate.true_ A⌉ := by 103 | rw [prod.map_map, prod.map_map] 104 | ext; simp 105 | rw [prod.map_snd, prod.map_snd, hom_comm] 106 | have h_terminal : 107 | (prod.map (𝟙 A) (terminal.from (hom A B)) ≫ prod.fst) ≫ terminal.from A = terminal.from _ := 108 | Unique.eq_default _ 109 | rw [assoc, comm_middle, ←assoc, comm_left, assoc, Predicate.NameDef] 110 | dsimp [Predicate.true_] 111 | rw [←assoc, ←assoc, h_terminal] 112 | 113 | /-- The evaluation map eval : A ⨯ B^A ⟶ B. -/ 114 | def eval (A B : C) : A ⨯ (hom A B) ⟶ B := 115 | IsPullback.lift (isPullback _) _ _ (eval_def_comm _ _) 116 | 117 | /-- This states the commutativity of the square relating 118 | `eval A B` to `singleton B` and `homToGraph A B`, which 119 | arises from its definition. 120 | -/ 121 | lemma eval_condition (A B : C) : 122 | eval A B ≫ singleton B 123 | = prod.map (𝟙 _) (homToGraph A B) ≫ ((prod.associator _ _ _).inv ≫ in_ (B ⨯ A))^ := 124 | IsPullback.lift_fst (isPullback _) _ _ (eval_def_comm _ _) 125 | 126 | variable {A B X : C} (f : A ⨯ X ⟶ B) 127 | 128 | /-- Useful definition in the context of the construction of `eval A B`. 129 | This is the composition of `homMap f` with `homToGraph A B`, as exhibited 130 | in `homMapCondition` below. 131 | -/ 132 | abbrev h_map : X ⟶ pow (B ⨯ A) := 133 | ((prod.associator _ _ _).hom ≫ prod.map (𝟙 _) f ≫ Predicate.eq _)^ 134 | 135 | 136 | /-- Helper lemma which shows that the appropriate square commutes 137 | in the definition of `homMap`. 138 | -/ 139 | lemma homMap_comm : 140 | h_map f ≫ (((prod.associator B A (pow (B ⨯ A))).inv 141 | ≫ in_ (B ⨯ A))^ ≫ Predicate.isSingleton B)^ = 142 | terminal.from X ≫ ⌈Predicate.true_ A⌉ := by 143 | -- consider (1⨯f) ≫ (eq B) : B ⨯ A ⨯ X ⟶ Ω C. 144 | let id_f'eq : B ⨯ A ⨯ X ⟶ Ω C := prod.map (𝟙 _) f ≫ Predicate.eq _ 145 | -- h is the map that, in `Set`, takes an element of X to the graph of the corresponding function. 146 | -- We want to lift this to a map X ⟶ Exp A B. 147 | -- The idea is to show that this map actually "maps elements of X to graphs of functions", which, 148 | -- in an arbitrary topos, is the same as checking commutativity of the obvious square. 149 | let h : X ⟶ pow (B ⨯ A) := (((prod.associator _ _ _).hom ≫ id_f'eq)^) 150 | -- h is by definition a P-transpose 151 | have h_condition : (prod.associator _ _ _).hom ≫ id_f'eq 152 | = (prod.map (prod.map (𝟙 _) (𝟙 _)) h) ≫ in_ _ := by 153 | rw [prod.map_id_id] 154 | symm 155 | apply Power.comm 156 | -- moving the associator to the rhs of `h_condition`. 157 | have h_condition₂ : id_f'eq 158 | = (prod.associator _ _ _).inv ≫ (prod.map (prod.map (𝟙 _) (𝟙 _)) h) ≫ in_ _ := by 159 | rw [←h_condition, ←assoc, (prod.associator _ _ _).inv_hom_id, id_comp] 160 | -- this is the map v: A ⨯ P(B⨯A) ⟶ P(B) which was used in the definition of `Exp A B`. 161 | let v : A ⨯ pow (B ⨯ A) ⟶ pow B := (((prod.associator _ _ _).inv ≫ in_ (B ⨯ A))^) 162 | -- v is by definition a P-transpose 163 | have v_condition : (prod.associator _ _ _).inv ≫ in_ (B ⨯ A) 164 | = prod.map (𝟙 _) v ≫ in_ _ := (comm _ _).symm 165 | have lhs : (prod.map (𝟙 A) h ≫ v ≫ Predicate.isSingleton B)^ 166 | = h ≫ (v ≫ Predicate.isSingleton B)^ := by 167 | apply Power.unique 168 | rw [prod.map_id_comp, assoc _ _ (in_ A), Power.comm, ←assoc] 169 | rw [←lhs] 170 | -- Claim that f ≫ {•}_B = (1⨯h) ≫ v. 171 | -- This is obtained by showing that both maps are the P-transpose of (1⨯f) ≫ (eq B). 172 | -- There might be a slightly faster way to do this. 173 | have transpose₁ : id_f'eq^ = f ≫ singleton _ := by 174 | apply Power.unique 175 | dsimp only [Topos.singleton] 176 | rw [prod.map_id_comp, assoc, (comm B (Predicate.eq B))] 177 | have shuffle_h_around : (prod.associator B A X).inv ≫ (prod.map (prod.map (𝟙 _) (𝟙 _)) h) 178 | = prod.map (𝟙 _) (prod.map (𝟙 _) h) ≫ (prod.associator _ _ _).inv := by simp 179 | have transpose₂ : id_f'eq^ = (prod.map (𝟙 _) h) ≫ v := by 180 | apply Power.unique 181 | rw [h_condition₂, ←assoc, shuffle_h_around, prod.map_id_comp, assoc _ _ (in_ B), 182 | ←v_condition, assoc] 183 | have eqn₁ : f ≫ singleton _ = (prod.map (𝟙 _) h) ≫ v := transpose₁.symm.trans transpose₂ 184 | -- now compose by the `isSingleton B` predicate. 185 | have eqn₂ : f ≫ singleton _ ≫ Predicate.isSingleton _ 186 | = (prod.map (𝟙 _) h) ≫ v ≫ Predicate.isSingleton _ := by 187 | rw [←assoc, ←assoc, eqn₁] 188 | rw [←eqn₂] 189 | -- from here, the argument is mostly definition unpacking. 190 | apply Power.unique 191 | dsimp only [name, Predicate.true_, Predicate.isSingleton] 192 | have f_terminal : f ≫ terminal.from B = terminal.from _ := Unique.eq_default _ 193 | have rightUnitor_terminal : (Limits.prod.rightUnitor A).hom ≫ terminal.from A 194 | = terminal.from _ := 195 | Unique.eq_default _ 196 | have A_X_terminal : prod.map (𝟙 A) (terminal.from X) ≫ terminal.from (A ⨯ ⊤_ C) 197 | = terminal.from _ := 198 | Unique.eq_default _ 199 | have obv : terminal.from (A ⨯ ⊤_ C) ≫ t C 200 | = prod.map (𝟙 A) ((terminal.from (A ⨯ ⊤_ C) ≫ t C)^) ≫ in_ A := 201 | (comm _ _).symm 202 | have map_def : (Limits.prod.rightUnitor A).hom = prod.fst := rfl 203 | rw [MonoClassifier.comm (singleton _), ←assoc, ←map_def, rightUnitor_terminal, ←assoc, 204 | f_terminal, prod.map_id_comp, assoc, ←obv, ←assoc, A_X_terminal] 205 | 206 | /-- Given a map `f : A ⨯ X ⟶ B`, `homMap f` 207 | is the associated map `X ⟶ Hom A B`. 208 | -/ 209 | def homMap : X ⟶ hom A B := 210 | pullback.lift (h_map f) (terminal.from X) (homMap_comm f) 211 | 212 | /-- composing `homMap f` with the map sending a morphism to its graph 213 | is the map `h_map f` defined above. 214 | -/ 215 | @[simp] 216 | lemma homMap_condition : homMap f ≫ (homToGraph A B) = h_map f := 217 | pullback.lift_fst _ _ _ 218 | 219 | /-- `homMap f` satisfies the condition that 220 | the following diagram commutes: 221 | ``` 222 | A ⨯ X ---f----> B 223 | | ^ 224 | | / 225 | | / 226 | (𝟙 A) ⨯ homMap / 227 | | / 228 | | eval A B 229 | | / 230 | | / 231 | | / 232 | | / 233 | v / 234 | A ⨯ (Hom A B) 235 | ``` 236 | -/ 237 | @[reassoc] 238 | theorem hom_exponentiates : prod.map (𝟙 _ ) (homMap f) ≫ eval A B = f := by 239 | rw [←cancel_mono (singleton B), assoc, eval_condition, ←assoc, 240 | prod.map_map, id_comp, homMap_condition] 241 | have h : transposeInv (f ≫ singleton B) 242 | = transposeInv (prod.map (𝟙 A) (h_map f) 243 | ≫ transpose ((prod.associator B A (Power.pow (B ⨯ A))).inv ≫ in_ (B ⨯ A))) := by 244 | rw [transposeInv, transposeInv, prod.map_id_comp, assoc, singleton, 245 | Power.comm, prod.map_id_comp, assoc, Power.comm, ←assoc] 246 | have h' : (prod.map (𝟙 B) (prod.map (𝟙 A) (h_map f)) 247 | ≫ (prod.associator B A (Power.pow (B ⨯ A))).inv) 248 | = (prod.associator B A X).inv ≫ (prod.map (𝟙 _) (h_map f)) := by simp 249 | rw [h', assoc, h_map, Power.comm, ←assoc, Iso.inv_hom_id, id_comp] 250 | have h₀ := congrArg (fun k => transpose k) h 251 | have t₁ : ((f ≫ singleton B)^)^ = f ≫ singleton B := (transposeEquiv _ _).right_inv _ 252 | have t₂ : (((prod.map (𝟙 A) (h_map f) 253 | ≫ ((prod.associator B A (Power.pow (B ⨯ A))).inv ≫ in_ (B ⨯ A))^))^)^ 254 | = (prod.map (𝟙 A) (h_map f) ≫ ((prod.associator B A (Power.pow (B ⨯ A))).inv ≫ in_ (B ⨯ A))^) := 255 | transpose_right_inv _ 256 | simp only [t₁, t₂] at h₀ 257 | exact h₀.symm 258 | 259 | /-- Proof of the uniqueness of `homMap f` as a map 260 | `X ⟶ Hom A B` satisfying commutativity of the associated 261 | diagram. See `Hom_Exponentiates`. 262 | -/ 263 | theorem Hom_Unique {exp' : X ⟶ hom A B} (h : prod.map (𝟙 _) exp' ≫ (eval A B) = f) : 264 | homMap f = exp' := by 265 | have h_singleton := congrArg (fun k ↦ k ≫ singleton B) h 266 | simp only at h_singleton 267 | let v : A ⨯ pow (B ⨯ A) ⟶ pow B := (((prod.associator _ _ _).inv ≫ in_ (B ⨯ A))^) 268 | -- want to rewrite (1⨯g) ≫ eval A B ≫ singleton B = (1⨯(g≫m)) ≫ v 269 | have rhs : eval A B ≫ singleton B = prod.map (𝟙 _) (homToGraph A B) ≫ v := by 270 | apply PullbackCone.IsLimit.lift_fst 271 | rw [assoc, rhs, ←assoc, ←prod.map_id_comp] at h_singleton 272 | let id_f'eq : B ⨯ A ⨯ X ⟶ Ω C := prod.map (𝟙 _) f ≫ Predicate.eq _ 273 | have h₁ : id_f'eq^ = f ≫ singleton B := by 274 | apply Power.unique 275 | dsimp only [id_f'eq, singleton] 276 | rw [prod.map_id_comp, assoc, comm _ (Predicate.eq B)] 277 | have h₂ : (prod.map (𝟙 _) (prod.map (𝟙 _) (exp' ≫ homToGraph A B)) 278 | ≫ (prod.associator _ _ _).inv ≫ in_ (B ⨯ A))^ 279 | = prod.map (𝟙 _) (exp' ≫ homToGraph A B) ≫ v := by 280 | apply Power.unique 281 | rw [prod.map_id_comp, assoc, Power.comm] 282 | have h₃ := Power.comm _ ((prod.map (𝟙 B) (prod.map (𝟙 A) (exp' ≫ homToGraph A B)) 283 | ≫ (prod.associator B A (Power.pow (B ⨯ A))).inv ≫ in_ (B ⨯ A))) 284 | rw [h₂, h_singleton, ←h₁, Power.comm _ id_f'eq, ←assoc] at h₃ 285 | have h' := hom_exponentiates f 286 | have h'_singleton := congrArg (fun k ↦ k ≫ singleton B) h' 287 | simp only at h'_singleton 288 | rw [assoc, rhs, ←assoc, ←prod.map_id_comp] at h'_singleton 289 | have h₂' : (prod.map (𝟙 _) (prod.map (𝟙 _) (homMap f ≫ homToGraph A B)) 290 | ≫ (prod.associator _ _ _).inv ≫ in_ (B ⨯ A))^ 291 | = prod.map (𝟙 _) (homMap f ≫ homToGraph A B) ≫ v := by 292 | apply Power.unique 293 | rw [prod.map_id_comp, assoc, Power.comm] 294 | have h₃' := Power.comm _ ((prod.map (𝟙 B) (prod.map (𝟙 A) (homMap f ≫ homToGraph A B)) 295 | ≫ (prod.associator B A (Power.pow (B ⨯ A))).inv ≫ in_ (B ⨯ A))) 296 | rw [h₂', h'_singleton, ←h₁, Power.comm _ id_f'eq, ←assoc] at h₃' 297 | 298 | have hx := h₃'.symm.trans h₃ 299 | have c₀ : prod.map (𝟙 B) (prod.map (𝟙 A) (exp' ≫ homToGraph A B)) ≫ (prod.associator _ _ _).inv 300 | = (prod.associator _ _ _).inv ≫ (prod.map (𝟙 _) (exp' ≫ homToGraph A B)) := by simp 301 | have c₁ : prod.map (𝟙 B) (prod.map (𝟙 A) (homMap f ≫ homToGraph A B)) 302 | ≫ (prod.associator _ _ _).inv 303 | = (prod.associator _ _ _).inv ≫ (prod.map (𝟙 _) (homMap f ≫ homToGraph A B)) := by simp 304 | rw [c₀, c₁] at hx 305 | have hy := congrArg (fun k ↦ (prod.associator B A X).hom ≫ k) hx 306 | simp only at hy 307 | rw [←assoc, ←assoc, Iso.hom_inv_id, id_comp, ←assoc, ←assoc, Iso.hom_inv_id, id_comp] at hy 308 | have hz := congrArg (fun k ↦ transpose k) hy 309 | simp only at hz 310 | rw [transpose_right_inv, transpose_right_inv] at hz 311 | rw [cancel_mono] at hz 312 | assumption 313 | 314 | variable {Y Z : C} 315 | 316 | /-- The inverse to `Hom_map`, which sends a morphism `X ⟶ Hom Y Z` 317 | to its "un-curried" version `Y ⨯ X ⟶ Z`. 318 | -/ 319 | abbrev homMapInv (f : X ⟶ hom Y Z) : Y ⨯ X ⟶ Z := prod.map (𝟙 _) f ≫ eval _ _ 320 | 321 | /-- The equivalence between arrows `A ⨯ X ⟶ B` and arrows `X ⟶ Hom A B`. -/ 322 | def ExpAdjEquiv (A B X : C) : (A ⨯ X ⟶ B) ≃ (X ⟶ hom A B) where 323 | toFun := homMap 324 | invFun := homMapInv 325 | left_inv := fun f => hom_exponentiates f 326 | right_inv := by 327 | intro f 328 | apply Hom_Unique; rfl 329 | 330 | /-- The map `Hom A X ⟶ Hom A Y` associated to a map `X ⟶ Y`. 331 | This is how `ExpFunctor` acts on morphisms. 332 | -/ 333 | def ExpHom {X Y : C} (A : C) (f : X ⟶ Y) : hom A X ⟶ hom A Y := homMap (eval A _ ≫ f) 334 | 335 | /-- The covariant functor `C ⥤ C` associated to an object `A : C` 336 | sending an object `B` to the "internal hom" `Hom A B`. 337 | -/ 338 | def ExpFunctor (A : C) : C ⥤ C where 339 | obj := fun B ↦ hom A B 340 | map := fun {X Y} f ↦ ExpHom A f 341 | map_id := by 342 | intro X 343 | dsimp only [ExpHom] 344 | rw [comp_id] 345 | apply Hom_Unique 346 | rw [prod.map_id_id, id_comp] 347 | map_comp := by 348 | intro X Y Z f g 349 | change ExpHom A (f ≫ g) = ExpHom A f ≫ ExpHom A g 350 | dsimp only [ExpHom] 351 | apply Hom_Unique 352 | rw [prod.map_id_comp, assoc, hom_exponentiates, hom_exponentiates_assoc, assoc] 353 | 354 | 355 | /-- A topos is a monoidal category with monoidal structure coming from binary products. -/ 356 | instance ToposMonoidal : MonoidalCategory C := monoidalOfHasFiniteProducts C 357 | 358 | /-- The adjunction between the product and the "internal hom" `Hom A B`. -/ 359 | def TensorHomAdjunction (A : C) : MonoidalCategory.tensorLeft A ⊣ ExpFunctor A := by 360 | apply Adjunction.mkOfHomEquiv 361 | fapply Adjunction.CoreHomEquiv.mk 362 | 363 | · intro X B 364 | exact ExpAdjEquiv A B X 365 | 366 | · intro X X' Y f g 367 | change prod.map (𝟙 _) (f ≫ g) ≫ eval _ _ = (prod.map (𝟙 _) f) ≫ prod.map (𝟙 _) g ≫ eval _ _ 368 | rw [prod.map_map_assoc, id_comp] 369 | 370 | · intro X Y Y' f g 371 | change homMap (f ≫ g) = homMap f ≫ ExpHom A g 372 | apply Hom_Unique 373 | dsimp only [ExpHom] 374 | rw [prod.map_id_comp, assoc, hom_exponentiates, hom_exponentiates_assoc] 375 | 376 | end 377 | 378 | section 379 | 380 | variable [ChosenFiniteProducts C] [HasClassifier C] [HasPullbacks C] [HasPowerObjects C] 381 | 382 | 383 | -- instance exponentiable (B : C) : Exponentiable B := Exponentiable.mk B (ExpFunctor B) (by 384 | -- exact TensorHomAdjunction B) 385 | 386 | /- 387 | TODO: figure out what the right approach is to get the result below. 388 | 389 | The main issue is this: `ChosenFiniteProducts C` is required for `CartesianClosed C`. 390 | this is because often there is an explicit formulation for finite products. 391 | However, this conflicts with the current definitions in files leading up to the definition 392 | of a topos. For example, a subobject classifier is in terms of a map `⊤_ C ⟶ Ω`, not a map 393 | `𝟙_ C ⟶ Ω`, and it Power.lean everything is in terms of the `HasBinaryProducts` product `⨯` 394 | instead of the monoidal structure `⊗` product. 395 | -/ 396 | 397 | def cartesianClosed : CartesianClosed C := sorry 398 | 399 | end 400 | end CategoryTheory.Topos 401 | --------------------------------------------------------------------------------