├── .gitignore ├── 2ltt ├── .project ├── Makefile ├── facts.lean ├── fibrant.lean ├── fibrantlimits.lean ├── fibrantlimits_aux.lean ├── finite.lean ├── inverse.lean ├── limit.lean ├── matching.lean ├── matching_facts.lean ├── pullbacks.lean ├── simplicial.lean └── types │ ├── prod.lean │ └── sigma.lean ├── Dockerfile ├── Makefile ├── README.md ├── extra └── init.el └── release.sh /.gitignore: -------------------------------------------------------------------------------- 1 | *.clean 2 | *.olean 3 | *.d 4 | *.ilean 5 | *.ninja* 6 | build.ninja 7 | TAGS 8 | code.zip -------------------------------------------------------------------------------- /2ltt/.project: -------------------------------------------------------------------------------- 1 | # Lean project file 2 | 3 | # Include all .lean files under this directory 4 | + *.lean 5 | 6 | # Exclude flycheck generated temp files 7 | - flycheck*.lean 8 | 9 | # Exclude emacs temp files 10 | - .#*.lean -------------------------------------------------------------------------------- /2ltt/Makefile: -------------------------------------------------------------------------------- 1 | build : 2 | linja 3 | 4 | clean : 5 | rm ./*.clean ./*.olean ./*.d ./*.ilean 6 | rm ./types/*.clean ./types/*.olean ./types/*.d ./types/*.ilean 7 | 8 | default : build 9 | -------------------------------------------------------------------------------- /2ltt/facts.lean: -------------------------------------------------------------------------------- 1 | import data.equiv 2 | 3 | open equiv equiv.ops eq eq.ops 4 | 5 | -- some definitions and facts about equality and equivalence 6 | -- which we could not find in standard library 7 | 8 | namespace eq 9 | 10 | definition ap {A B : Type}(f : A → B){a a' : A} 11 | (p : a = a') : f a = f a' := p ▹ eq.refl _ 12 | 13 | definition apd {A : Type}{B : A → Type}(f : Π (a : A), B a) 14 | {a a' : A}(p : a = a') : p ▹ f a = f a' := 15 | eq.drec (eq.refl _) p 16 | 17 | definition naturality_subst {X Y : Type}{x x' : X}{P : Y → Type} 18 | (f : X → Y)(p : x = x')(u : P (f x)) : ap f p ▹ u = p ▹ u := 19 | eq.drec (eq.refl _) p 20 | 21 | definition transport_concat [simp] {A : Type } {a b c: A} (P : A → Type) 22 | (p : a = b) (q : b = c) (u : P a) : q ▹ (p ▹ u) = p ⬝ q ▹ u := begin cases p, cases q, reflexivity end 23 | 24 | definition concat_inv {A : Type } {a b : A} {P : A → Type} 25 | (p : a = b) (u : P a) : #eq.ops p ⬝ p⁻¹ ▹ u = u := by cases p; reflexivity 26 | 27 | end eq 28 | 29 | open eq 30 | 31 | namespace equiv 32 | 33 | open eq.ops sigma.ops function 34 | 35 | variables {A A' : Type} 36 | 37 | definition pi_congr₁ [instance] {F' : A' → Type} [φ : A ≃ A'] 38 | : (Π (a : A), F' (φ ∙ a)) ≃ (Π (a : A'), F' a) := 39 | match φ with mk f g l r := 40 | mk (λ k x', r x' ▹ k (g x')) 41 | (λ h x, h (f x)) 42 | (λ k, funext (λ x, 43 | calc r (f x) ▹ k (g (f x)) 44 | = ap f (l x) ▹ k (g (f x)) : proof_irrel (r (f x)) (ap f (l x)) 45 | ... = l x ▹ k (g (f x)) : naturality_subst f (l x) (k (g (f x))) 46 | ... = k x : apd k (l x))) 47 | (λ h, funext (λ x', apd h (r x'))) 48 | end 49 | 50 | definition pi_congr₂ [instance] {F G : A → Type} [φ : Π a, F a ≃ G a] 51 | : (Π (a : A), F a) ≃ (Π (a : A), G a) := 52 | equiv.mk 53 | (λ x a, (φ a) ∙ x a) 54 | (λ x a, inv (x a)) 55 | (λ x, funext (λ a, (left_inv _ (G a) _) ▹ rfl)) 56 | (λ x, funext (λ a, (right_inv (F a) _ _) ▹ rfl)) 57 | 58 | definition pi_congr [instance] {F : A → Type} {F' : A' → Type} [φ : A ≃ A'] [φ' : Π a, F a ≃ F' (φ ∙ a)] := 59 | equiv.trans (@pi_congr₂ _ _ _ φ') (@pi_congr₁ _ _ _ φ) 60 | 61 | definition sigma_eq_congr {A: Type} {F : A → Type} {a a': A} {b : F a} {b' : F a'} : 62 | (Σ (p : a = a'), ((p ▹ b) = b')) → ⟨ a, b ⟩ = ⟨a', b'⟩ := 63 | begin intro p, cases p with [p₁, p₂], cases p₁, cases p₂, esimp end 64 | 65 | definition sigma_congr₁ [instance] {A B: Type} {F : B → Type} [φ : A ≃ B]: 66 | (Σ a : A, F (φ ∙ a)) ≃ Σ b : B, F b := 67 | equiv.mk 68 | (λ x , ⟨ φ ∙ x.1, x.2 ⟩ ) 69 | (λ x, ⟨ inv x.1, begin unfold fn, unfold inv, rewrite right_inv, apply x.2 end⟩ ) 70 | (λ x, begin 71 | unfold fn at *, 72 | cases φ with [f, g, l, r], esimp at *, unfold fn, 73 | cases x with [x₁, x₂], 74 | unfold function.right_inverse at *, unfold function.left_inverse at *, unfold inv, esimp, 75 | apply sigma_eq_congr, refine ⟨l x₁,_⟩, 76 | calc 77 | (l x₁ ▹ ((r (f x₁))⁻¹ ▹ x₂)) 78 | = (l x₁ ▹ (eq.symm (ap f (l x₁)) ▹ x₂)) : rfl 79 | ... = (l x₁ ▹ (eq.symm (l x₁) ▹ x₂)) : naturality_subst 80 | ... = ((l x₁ ⬝ (eq.symm (l x₁))) ▹ x₂) : transport_concat 81 | ... = x₂ : concat_inv (l x₁) 82 | end) 83 | begin 84 | intro x, cases x with [x₁, x₂], 85 | cases φ with [f, g, l, r], unfold function.right_inverse at *, unfold function.left_inverse at *, esimp at *, 86 | apply sigma_eq_congr, refine ⟨_,_⟩, apply r, 87 | calc 88 | #eq.ops 89 | r x₁ ▹ (r x₁)⁻¹ ▹ x₂ 90 | = r x₁ ⬝ (r x₁)⁻¹ ▹ x₂ : transport_concat 91 | ... = x₂ : concat_inv (r x₁) 92 | end 93 | 94 | definition sigma_congr₂ [instance] {A : Type} {F G : A → Type} [φ : Π a : A, F a ≃ G a] : 95 | (Σ a, F a) ≃ Σ a, G a := 96 | equiv.mk 97 | (λ x, ⟨x.1, (φ x.1) ∙ x.2⟩) 98 | (λ x, ⟨x.1, inv x.2⟩) 99 | begin 100 | unfold left_inverse, intros, 101 | cases x with [p₁, p₂], esimp, 102 | congruence, unfold fn, unfold inv, 103 | rewrite left_inv 104 | end 105 | begin 106 | unfold right_inverse, unfold left_inverse, intros, 107 | cases x with [p₁, p₂], congruence, unfold fn, unfold inv, 108 | rewrite right_inv 109 | end 110 | 111 | definition sigma_congr {A B : Type} {F : A → Type} {G : B → Type} 112 | [φ : A ≃ B] [φ' : Π a : A, F a ≃ G (φ ∙ a)] : 113 | (Σ a, F a) ≃ Σ a, G a := equiv.trans sigma_congr₂ sigma_congr₁ 114 | 115 | definition sigma_swap {A B : Type} {F : A → B → Type} : 116 | (Σ (a : A) (b : B), F a b) ≃ Σ (b : B) (a : A), F a b := 117 | equiv.mk (λ x, ⟨x.2.1,⟨x.1,x.2.2⟩⟩) 118 | (λ x, ⟨x.2.1,⟨x.1,x.2.2⟩⟩) 119 | (λ x, by cases x with [p1,p2]; cases p2; reflexivity) 120 | (λ x, by cases x with [p1,p2]; cases p2; reflexivity) 121 | 122 | definition iff_impl_equiv {A B : Prop} (Hiff : A ↔ B) : A ≃ B := 123 | match Hiff with 124 | | iff.intro f g := equiv.mk f g (λx, proof_irrel _ _) (λx, proof_irrel _ _) 125 | end 126 | 127 | end equiv 128 | 129 | namespace function 130 | definition happly {A B : Type} {f g : A → B} : f = g -> ∀ x, f x = g x := 131 | begin 132 | intros H x, rewrite H 133 | end 134 | end function 135 | -------------------------------------------------------------------------------- /2ltt/fibrant.lean: -------------------------------------------------------------------------------- 1 | import data.equiv .facts 2 | 3 | open eq 4 | 5 | -- we use the isomorphism definition from the standard library 6 | notation X `≃` Y := equiv X Y 7 | constant is_fibrant' : Type → Prop 8 | 9 | structure is_fibrant [class] (X : Type) := mk :: 10 | fib_internal : is_fibrant' X 11 | 12 | constant equiv_is_fibrant {X Y : Type}(e : X ≃ Y)[fib : is_fibrant X] : is_fibrant Y 13 | 14 | constant unit_is_fibrant' : is_fibrant' unit 15 | constant polyunit_is_fibrant' : is_fibrant' poly_unit 16 | 17 | constant pi_is_fibrant' {X : Type}{Y : X → Type} 18 | : is_fibrant X 19 | → (Π (x : X), is_fibrant (Y x)) 20 | → is_fibrant' (Π (x : X), Y x) 21 | 22 | constant sigma_is_fibrant' {X : Type}{Y : X → Type} 23 | : is_fibrant X 24 | → (Π (x : X), is_fibrant (Y x)) 25 | → is_fibrant' (Σ (x : X), Y x) 26 | 27 | constant prod_is_fibrant' {X Y : Type} 28 | : is_fibrant X -> is_fibrant Y -> is_fibrant' (X × Y) 29 | 30 | -- explicit universe level declaration allows to use tactics; 31 | -- otherwise, proofs fail to unify level variables sometime 32 | constant fib_eq.{u} {X : Type.{u}}[is_fibrant X] : X → X → Type.{u} 33 | 34 | -- the inner equality (or path-equality) type (in the sense of HoTT) 35 | namespace fib_eq 36 | infix ` =ᵒ `:28 := fib_eq 37 | constant reflo {X : Type}[is_fibrant X](x : X) : x =ᵒ x 38 | notation ` reflᵒ ` := reflo 39 | 40 | constant elimo {X : Type}[is_fibrant X]{x : X}{P : Π (y : X), x =ᵒ y → Type} 41 | [Π (y : X)(p : x =ᵒ y), is_fibrant (P y p)] 42 | (d : P x (reflᵒ x)) : Π (y : X)(p : x =ᵒ y), P y p 43 | notation ` elimᵒ ` := elimo 44 | 45 | constant elimo_β {X : Type}[is_fibrant X]{x : X}{P : Π (y : X), x =ᵒ y → Type} 46 | [Π (y : X)(p : x =ᵒ y), is_fibrant (P y p)] 47 | (d : P x (reflᵒ x)) : elimᵒ d x (reflᵒ x) = d 48 | 49 | notation ` elim_βᵒ ` := elimo_β 50 | 51 | definition idp [reducible] [constructor] {X : Type}[is_fibrant X] {a : X} := reflo a 52 | 53 | attribute elimo_β [simp] 54 | attribute elimo [recursor] 55 | end fib_eq 56 | 57 | constant fib_eq_is_fibrant' {X : Type}[is_fibrant X](x y : X) : is_fibrant' (fib_eq x y) 58 | 59 | structure Fib : Type := mk :: 60 | (pretype : Type) 61 | (fib : is_fibrant pretype) 62 | 63 | attribute Fib.pretype [coercion] 64 | attribute Fib.fib [instance] 65 | 66 | constant fib_is_fibrant' : is_fibrant' Fib 67 | 68 | -- instances 69 | 70 | definition unit_is_fibrant [instance] : is_fibrant unit := is_fibrant.mk unit_is_fibrant' 71 | definition polyunit_is_fibrant [instance] : is_fibrant poly_unit := is_fibrant.mk polyunit_is_fibrant' 72 | 73 | definition pi_is_fibrant [instance] {X : Type}{Y : X → Type} 74 | [fibX : is_fibrant X] [fibY : Π (x : X), is_fibrant (Y x)] : 75 | is_fibrant (Π (x : X), Y x) := 76 | is_fibrant.mk (pi_is_fibrant' fibX fibY) 77 | 78 | definition sigma_is_fibrant [instance] {X : Type}{Y : X → Type} 79 | [fibX : is_fibrant X] [fibY : Π (x : X), is_fibrant (Y x)] : 80 | is_fibrant (Σ (x : X), Y x) := 81 | is_fibrant.mk (sigma_is_fibrant' fibX fibY) 82 | 83 | definition prod_is_fibrant [instance] {X Y : Type} 84 | [fibX : is_fibrant X] [fibY : is_fibrant Y] : 85 | is_fibrant (X × Y) := is_fibrant.mk (prod_is_fibrant' fibX fibY) 86 | 87 | definition fib_is_fibrant [instance] : is_fibrant Fib := is_fibrant.mk fib_is_fibrant' 88 | 89 | definition fib_eq_is_fibrant [instance] {X : Type}[is_fibrant X](x y : X) : 90 | is_fibrant (fib_eq x y) := 91 | is_fibrant.mk (fib_eq_is_fibrant' x y) 92 | 93 | -- basics 94 | 95 | open prod 96 | 97 | namespace fib_eq 98 | open function 99 | 100 | variables {X: Fib} 101 | variables {Y: Fib} 102 | variables {Z: Fib} 103 | 104 | attribute reflo [refl] 105 | definition symm [symm] {x y : X} : x =ᵒ y → y =ᵒ x := 106 | elimᵒ (reflᵒ x) y 107 | definition symm_β [simp] {x : X} : symm (reflᵒ x) = reflᵒ x := 108 | elim_βᵒ (reflᵒ x) 109 | definition trans [trans] {x y z : X} : x =ᵒ y → y =ᵒ z → x =ᵒ z := λ p, 110 | elimᵒ p z 111 | 112 | definition trans_β [simp] {x y : X}(p : x =ᵒ y) : trans p (reflᵒ y) = p := 113 | elim_βᵒ p 114 | 115 | definition trans_β' [simp] {x : X} : trans (reflᵒ x) (reflᵒ x) = reflᵒ x := by simp 116 | 117 | definition trans' {x y z : X} (p : x =ᵒ y) (q : y =ᵒ z) : x =ᵒ z := 118 | (elimᵒ (elimᵒ idp z) y) p q 119 | 120 | -- Alternative proof of transitivity using tactics. 121 | definition trans'' {x y z : X} (p: x =ᵒ y) (q : y =ᵒ z) : x =ᵒ z := 122 | by induction p using elimo; exact q 123 | 124 | definition transport [subst] {x y : X}{P : X → Type}[Π (x : X), is_fibrant (P x)] 125 | (p : x =ᵒ y)(d : P x) : P y := 126 | elimᵒ d y p 127 | definition transport_β [simp] {x : X}{P : X → Type}[Π (x : X), is_fibrant (P x)](d : P x) : 128 | transport (reflᵒ x) d = d := 129 | elim_βᵒ d 130 | 131 | notation p ▹ d := transport p d 132 | 133 | infixl ⬝ := trans 134 | postfix ⁻¹ := symm 135 | 136 | definition symm_trans {x y : X} (p : x =ᵒ y) : p⁻¹ ⬝ p =ᵒ reflᵒ y := 137 | by induction p using elimo; simp 138 | 139 | definition symm_trans_β {x : X} : (reflᵒ x)⁻¹ ⬝ (reflᵒ x) = reflᵒ x := by simp 140 | 141 | definition assoc_trans {x y z t: X} (p : x =ᵒ y) (q : y =ᵒ z) (r : z =ᵒ t) : 142 | (p ⬝ (q ⬝ r)) =ᵒ ((p ⬝ q) ⬝ r) := 143 | by induction r using elimo; simp 144 | 145 | definition transport_trans {A : Fib } {a b c: A} {P : A → Fib} 146 | (p : a =ᵒ b) (q : b =ᵒ c) (u : P a) : #fib_eq q ▹ (p ▹ u) =ᵒ p ⬝ q ▹ u := 147 | by induction p using elimo; induction q using elimo; simp 148 | 149 | definition eq_transport_l {a₁ a₂ a₃ : X} (p : a₁ =ᵒ a₂) (q : a₁ =ᵒ a₃) : 150 | p ▹ q =ᵒ p⁻¹ ⬝ q := 151 | by induction p using elimo; induction q using elimo; simp 152 | 153 | definition eq_transport_r {a₁ a₂ a₃ : X} (p : a₂ =ᵒ a₃) (q : a₁ =ᵒ a₂) : 154 | p ▹ q =ᵒ q ⬝ p := 155 | by induction p using elimo; induction q using elimo; simp 156 | 157 | definition to_fib_eq { x y : X } : x = y -> x =ᵒ y := eq.rec (reflᵒ _) 158 | 159 | namespace ap 160 | -- action on paths 161 | 162 | -- notation for the transport along the strict equality 163 | notation p `▹s` x := eq.rec_on p x 164 | 165 | 166 | definition ap {x y : X} (f : X -> Y) : x =ᵒ y -> f x =ᵒ f y := 167 | elimᵒ (reflᵒ _) _ 168 | 169 | definition ap_β [simp] {x : X} (f : X -> Y) : ap f (reflᵒ x) = reflᵒ (f x) := elim_βᵒ (reflᵒ (f x)) 170 | 171 | definition ap_trans {x y z : X} (f : X → Y) (p : x =ᵒ y) (q : y =ᵒ z) : 172 | ap f (p ⬝ q) =ᵒ (ap f p) ⬝ (ap f q) := by induction p using elimo; induction q using elimo; simp 173 | 174 | definition ap_symm {x y : X} (f : X → Y) (p : x =ᵒ y) : ap f p⁻¹ =ᵒ (ap f p)⁻¹ := 175 | by induction p using elimo; simp 176 | 177 | definition ap_compose {x y : X} (f : X → Y) (g : Y → Z) (p : x =ᵒ y) : ap g (ap f p) =ᵒ ap (g ∘ f) p := 178 | by induction p using elimo; simp 179 | 180 | definition ap_id {x y : X} (p : x =ᵒ y) : ap id p =ᵒ p := by induction p using elimo; simp 181 | 182 | definition ap₂ [reducible] {x x' : X} {y y' : Y} (f : X -> Y -> Z) 183 | (p : x =ᵒ x') (q : y =ᵒ y') : f x y =ᵒ f x' y' := (ap (λ x, f x y) p) ⬝ (ap (f x') q) 184 | 185 | definition ap₂_β {x : X} {y : Y} (f : X -> Y -> Z) : ap₂ _ (reflᵒ x) (reflᵒ y) = reflᵒ (f x y) := 186 | by simp 187 | 188 | definition apd {P : X → Fib} (f : Π x, P x) {x y : X} (p : x =ᵒ y) : p ▹ f x =ᵒ f y := 189 | by induction p using elimo; rewrite transport_β 190 | 191 | definition apd_β {P : X → Fib} (f : Π x, P x) {x y : X} : 192 | apd f (reflᵒ x) = (eq.symm (transport_β _) ▹s reflᵒ (f x)) := elim_βᵒ _ 193 | 194 | definition apd_β' {P : X → Fib} (f : Π x, P x) {x y : X} : 195 | ((transport_β (f x)) ▹s (apd f (reflᵒ x))) = reflᵒ (f x) := 196 | by unfold apd; rewrite elimo_β; rewrite eq.transport_concat 197 | 198 | end ap 199 | 200 | end fib_eq 201 | 202 | open fib_eq 203 | 204 | variables {A: Fib} 205 | variables {B: Fib} 206 | variables {C: Fib} 207 | 208 | 209 | structure is_contr (X : Type)[is_fibrant X] := mk :: 210 | (center : X) 211 | (contraction : Π (x : X), center =ᵒ x) 212 | 213 | open sigma.ops is_contr 214 | 215 | definition is_contr_equiv [instance] {X : Type}[is_fibrant X] : 216 | (Σ (c : X), Π (x : X), c =ᵒ x) ≃ is_contr X := 217 | equiv.mk 218 | (λ a, is_contr.mk a.1 (λx, a.2 x)) 219 | (λ a, ⟨center a,(λx, contraction a x)⟩) 220 | (λ a, by cases a; reflexivity) 221 | (λ a, by cases a; reflexivity) 222 | 223 | definition is_contr_is_fibrant [instance] (X : Type)[is_fibrant X] : is_fibrant (is_contr X) := 224 | equiv_is_fibrant is_contr_equiv 225 | 226 | definition is_trunc : Π (n : ℕ)(X : Type) [is_fibrant X], Type 227 | | @is_trunc 0 X fib := is_contr X 228 | | @is_trunc (nat.succ n) X fib := Π (x y : X), is_trunc n (x =ᵒ y) 229 | 230 | definition is_prop (X : Type) [is_fibrant X] := Π (x y : X), x =ᵒ y 231 | 232 | section truncated_types 233 | variables (X : Fib) 234 | 235 | definition is_contr_is_trunc : 236 | is_contr X → is_trunc 1 X := 237 | assume c : is_contr X, 238 | let path (x y : X) : x =ᵒ y := 239 | symm (is_contr.contraction c x) ⬝ is_contr.contraction c y in 240 | let l (x y : X)(p : x =ᵒ y) : path x y =ᵒ p := 241 | elimᵒ (!symm_trans) y p in 242 | λ x y, is_contr.mk (path x y) (l x y) 243 | 244 | definition inhab_is_contr_is_prop : (X → is_contr X) → is_trunc 1 X := 245 | assume H x, 246 | have is_contr X, from H x, 247 | have is_trunc 1 X, from is_contr_is_trunc X this, 248 | this x 249 | 250 | definition is_prop_is_trunc : is_prop X → is_trunc 1 X := 251 | assume p : is_prop X, 252 | have X → is_contr X, from λ x, is_contr.mk x (p x), 253 | show is_trunc 1 X, from inhab_is_contr_is_prop X this 254 | 255 | definition is_trunc_is_prop : is_trunc 1 X → is_prop X := λ t x y, 256 | is_contr.center (t x y) 257 | end truncated_types 258 | 259 | section singleton 260 | variables {X : Fib} 261 | definition singleton [reducible] (x : X) := Σ (y : X), y =ᵒ x 262 | definition singleton_contr (x : X) : is_contr (singleton x) := 263 | let l (y : X)(p : y =ᵒ x) : ⟨x , reflᵒ x⟩ =ᵒ ⟨y, p⟩ := elimᵒ !reflᵒ x p in 264 | is_contr.mk ⟨x, reflᵒ x⟩ (λt, match t with ⟨y, p⟩ := l y p end) 265 | end singleton 266 | 267 | section fib_equivalences 268 | variables {X Y : Type}(f : X → Y)[is_fibrant X][is_fibrant Y] 269 | 270 | definition fibre [reducible] (y : Y) := Σ (x : X), f x =ᵒ y 271 | definition is_fib_equiv [reducible] [class] := Π (y : Y), is_contr (fibre f y) 272 | 273 | definition fib_equiv [reducible] (X Y : Type)[is_fibrant X][is_fibrant Y] : Type := 274 | Σ (f : X → Y), is_fib_equiv f 275 | 276 | definition id_is_fib_equiv {X : Type}[is_fibrant X] : is_fib_equiv (@id X) := singleton_contr 277 | 278 | notation X `≃` Y := fib_equiv X Y 279 | 280 | definition coerce {X Y : Fib} : X =ᵒ Y → X → Y := elimᵒ id Y 281 | definition coerce_is_fib_equiv [instance] {X Y : Fib}(p : X =ᵒ Y) : is_fib_equiv (coerce p) := 282 | begin 283 | induction p using elimo, 284 | unfold coerce, rewrite elimo_β, 285 | apply id_is_fib_equiv 286 | end 287 | 288 | definition coerce_fib_equiv {X Y : Fib}(p : X =ᵒ Y) : X ≃ Y := ⟨ coerce p, _ ⟩ 289 | 290 | end fib_equivalences 291 | 292 | 293 | definition Univalence := Π {X Y : Fib}, is_fib_equiv (@coerce_fib_equiv X Y) 294 | 295 | -- (strict) fibre 296 | definition fibreₛ [reducible] {X Y : Type} (f : X → Y) (y : Y) := Σ (x : X), f x = y 297 | 298 | open sigma.ops 299 | 300 | definition fibre_projection {X : Type}{Y : X → Type}(x : X) 301 | : fibreₛ (λ (p : Σ (x : X), Y x), p.1) x ≃ Y x 302 | := begin 303 | unfold fibreₛ, esimp, 304 | refine (equiv.mk _ _ _ _), 305 | { intro, cases a with [a₁, a₂], cases a₂, cases a₁, assumption}, 306 | { intros, apply (⟨⟨x,a⟩, rfl⟩) }, 307 | { unfold function.left_inverse, intro y, 308 | cases y with [y₁, y₂], cases y₁ with [z₁, z₂], esimp, 309 | esimp at *, congruence, cases y₂, congruence }, 310 | {apply (λ x, rfl)} 311 | end 312 | 313 | -- ref:def:fibration 314 | -- Definition 3.7 315 | definition is_fibration [reducible] {E B : Type} (p : E → B) := 316 | Π (b : B), is_fibrant (fibreₛ p b) 317 | 318 | definition is_fibration_alt.{u} {E B : Type.{max 1 u}} (p : E → B) := 319 | Σ (F : B → Fib.{u}), Π (b : B), F b ≃ fibreₛ p b 320 | -------------------------------------------------------------------------------- /2ltt/fibrantlimits.lean: -------------------------------------------------------------------------------- 1 | import .fibrant .fibrantlimits_aux .matching .inverse algebra.category 2 | .limit .pullbacks .finite 3 | import data.fin 4 | 5 | open invcat category functor matching_object Fib 6 | sigma.ops fincat natural_transformation eq 7 | 8 | open equiv equiv.ops 9 | 10 | namespace fiblimits 11 | universe variable u 12 | variables {C : Category.{1 1}} 13 | {D : Category} 14 | 15 | open subcat_obj reedy 16 | open sum nat fin function 17 | 18 | -- ref:thm:fibrant-limits 19 | -- Theorem 4.8 20 | definition fibrant_limit.{v} 21 | [invC : invcat C] [finC : is_obj_finite C] 22 | (X : C ⇒ U) (rfib : is_reedy_fibrant X) : 23 | is_fibrant (limit_obj (limit_in_pretype X)) := 24 | 25 | begin 26 | cases finC with [n, ψ], 27 | revert ψ, revert rfib, revert invC, revert X, revert C, 28 | -- induction on the cardinality of |C| 29 | induction n with [n', IHn], 30 | { intros C X invC rfib ψ, apply equiv_is_fibrant.{v} (@fincat_0_limit_unit_equiv _ ψ X)⁻¹ }, 31 | { intros C X invC rfib ψ, esimp, 32 | have inv_C : invcat C, from invC, 33 | cases invC, cases reflecting_id_ℕop with [φ, idrefl], 34 | -- choosing an element of C with the maximal rank 35 | have H : Σ z, (Π (y : C), φ y ≤ φ z), from @max_fun_to_ℕ _ φ _ ψ, 36 | cases H with [z, z_max_φ], 37 | 38 | -- removing z from C and showing that the resulting category 39 | -- is still inverse and finite 40 | have invC' : invcat (C_without_z z), from C_without_z_invcat z, 41 | have finC' : C_without_z z ≃ₛ fin n', from @C_without_z_is_obj_finite _ _ _ _, 42 | 43 | 44 | -- using equivalences 45 | apply equiv_is_fibrant, 46 | apply equiv.symm nat_unit_sigma_equiv, 47 | 48 | let q := matching_obj_map X z, 49 | have fibration_q : (is_fibration q), from rfib z, 50 | 51 | let p := map_L_to_Mz_alt z X, 52 | apply equiv_is_fibrant, apply equiv.symm, 53 | 54 | let C' := C_without_z z, 55 | let X' := Functor_from_C' z X, 56 | let limX' := lim_restricted X z, 57 | 58 | calc 59 | (Σ (c : Π y, X y), Π y y' f, X f (c y) = c y') 60 | -- split the limit in two pieces 61 | ≃ₛ (Σ c_z c, 62 | (Π (y : C') (f : z ⟶ obj y ), X f c_z = c y) × 63 | (Π (y y' : C') (f : @hom (subcat_obj _ _) _ y y'), X' f (c y) = c y')) : limit_two_piece_limit_equiv ψ idrefl z_max_φ 64 | /-- get a pullback of the span 65 | [lim X' --p--> matching_object M Z <<--q-- X z] 66 | where lim X' is the limit of X restricted to C' 67 | (so, L is Nat(𝟙,X')) -/ 68 | ... ≃ₛ (Σ c_z d, p d = q c_z) : two_piece_limit_pullback_p_q_equiv 69 | ... ≃ₛ (Σ d c_z, p d = q c_z) : equiv.sigma_swap 70 | ... ≃ₛ (Σ d c_z, q c_z = p d) : by apply @sigma_congr₂; intros; 71 | apply @sigma_congr₂; intros; 72 | apply (iff_impl_equiv (iff.intro eq.symm eq.symm)), 73 | 74 | -- to show that this pullback is fibrant we use facts that q is a fibration (from Reedy fibrancy of X) and 75 | -- that limX' is fibrant (from IH) 76 | have rfibX' : is_reedy_fibrant X', 77 | from @Functor_from_C'_reedy_fibrant _ z X inv_C φ z_max_φ idrefl rfib ⟨_,ψ⟩, 78 | have isFibLimX': is_fibrant limX', 79 | from @equiv_is_fibrant _ _ nat_unit_sigma_equiv (IHn _ _ _ _ _), 80 | refine @fibration_domain_is_fibrant _ (Fib.mk _ isFibLimX') (λpb, pb.1) _, 81 | refine Pullback'_is_fibrant q p, apply fibration_q 82 | } 83 | end 84 | end fiblimits 85 | -------------------------------------------------------------------------------- /2ltt/fibrantlimits_aux.lean: -------------------------------------------------------------------------------- 1 | import .fibrant .matching .inverse algebra.category 2 | .limit .pullbacks .finite .matching_facts 3 | import data.fin 4 | 5 | open invcat category functor matching_object Fib 6 | sigma.ops fincat natural_transformation eq 7 | 8 | universe variable u 9 | variables {C : Category.{1 1}} 10 | {D : Category} 11 | 12 | open subcat_obj reedy 13 | open sum nat fin function equiv 14 | 15 | definition max_fun_to_ℕ (φ : C → ℕ) {n : ℕ} [ψ : objects C ≃ₛ fin (nat.succ n)] : 16 | Σ z, (Π c, φ c ≤ φ z) := 17 | begin 18 | generalize φ, clear φ, generalize ψ, clear ψ, generalize C, clear C, 19 | induction n with [n', IHn], 20 | intros, 21 | { cases ψ with [f,g,l,r], esimp at *, 22 | existsi g (fin.zero 0), 23 | intro, assert H : f c = zero, apply eq_of_veq, cases (f c) with [v, Hlt], esimp, cases v, 24 | reflexivity, cases Hlt with [a, Hle], exfalso, 25 | rewrite -(succ_le_zero_iff_false (succ a)), apply Hle, rewrite -l, rewrite H }, 26 | { intros, 27 | have H : Σ z, @to_fun _ (fin (succ (nat.succ n'))) ψ z = maxi, from ⟨(inv_fun C maxi), !right_inv⟩, 28 | cases H with [z'', max_z], 29 | have ψ' : C_without_z z'' ≃ fin (succ n'), from (@C_without_z_is_obj_finite _ _ z'' _), 30 | have H : Π (φ : C_without_z z'' → ℕ), 31 | Σ (z : C_without_z z''), (Π c, φ c ≤ φ z), 32 | from IHn _ ψ', 33 | cases H (φ ∘ obj) with [z', Hmax], 34 | cases (@decidable_le (φ z'') (φ z')) with [z_le, z_ne_le], 35 | existsi z', intros, 36 | cases (@fincat.has_decidable_eq _ ⟨succ (succ n'), ψ⟩ c z'') with [c_eq_z, c_ne_z], 37 | { rewrite c_eq_z, assumption }, 38 | { apply Hmax (mk _ c_ne_z) }, 39 | existsi z'', intro, apply nat.le_of_eq_or_lt, 40 | cases (@fincat.has_decidable_eq _ ⟨succ (succ n'), ψ⟩ c z'') with [c_eq_z, c_ne_z], 41 | left, rewrite c_eq_z, right, 42 | have Hmax_c' : φ (obj (mk c c_ne_z)) ≤ φ (obj z'), from Hmax (mk c c_ne_z), esimp at *, 43 | have Hlt : φ z' < φ z'', from iff.elim_right (gt_iff_not_le _ _) z_ne_le, 44 | apply lt_of_le_of_lt Hmax_c' Hlt } 45 | end 46 | 47 | open invcat 48 | 49 | definition Functor_from_C'_eq (X : C ⇒ U) (z : C) (x' : C_without_z z) : 50 | X (@obj _ _ x') = (object (Functor_from_C' z X) x') := 51 | begin esimp end 52 | 53 | open poly_unit 54 | open reduced_coslice 55 | open reduced_coslice.red_coslice_obs 56 | 57 | -- for any object "a" from z//C (which is an arrow z⟶y with extra property) we show that codomain y cannot be z 58 | -- we show it using the definition of reduced coslice and the fact that C is an inverse category 59 | definition reduced_coslice_ne (z : C) (a : z//C) [invC : invcat C] : to a ≠ z := 60 | begin 61 | cases a with [y, f, non_id_f], esimp at *, 62 | intro p, cases p, 63 | apply non_id_f rfl, 64 | cases invC, cases reflecting_id_ℕop, 65 | apply (reflecting_id f ⟨rfl,rfl⟩).2 66 | end 67 | 68 | definition Functor_from_C'_reedy_fibrant (z : C) (X : C ⇒ U) [invC : invcat C] 69 | {φ : C ⇒ ℕop} {max_rank : ∀ x, φ x ≤ φ z} {reflecting_id : id_reflect φ} 70 | [rfibX : is_reedy_fibrant X] [finC : is_obj_finite C] 71 | : is_reedy_fibrant (Functor_from_C' z X) := 72 | begin 73 | let C' := C_without_z z, 74 | let X' := Functor_from_C' z X, 75 | unfold is_reedy_fibrant at *, 76 | unfold is_fibration at *, 77 | intros x b, 78 | let Hequiv := (@MO_equiv C z x φ reflecting_id max_rank _ _ X), 79 | let MO := @to_fun _ _ Hequiv b, 80 | assert isfibX' : is_fibrant (fibreₛ (matching_obj_map X (obj x)) (@to_fun _ _ Hequiv b)), begin apply (rfibX x MO) end, 81 | assert Hfeq': (fibreₛ (matching_obj_map X (obj x)) MO) ≃ₛ (fibreₛ (matching_obj_map (Functor_from_C' z X) x) b), 82 | begin 83 | unfold fibreₛ, apply @sigma_congr₂, intros, esimp, 84 | apply iff_impl_equiv, refine iff.intro _ _, 85 | { intros p, 86 | calc 87 | matching_obj_map X' x a 88 | = @inv_fun _ _ Hequiv (matching_obj_map X (obj x) a) : 89 | begin refine nat_trans_eq, apply funext, intro y, apply funext, intro uu, 90 | cases y, cases uu, esimp end 91 | ... = @inv_fun _ _ Hequiv (@to_fun _ _ Hequiv b) : ap _ p 92 | ... = b : @left_inv _ _ _ 93 | }, 94 | { intros p, 95 | calc 96 | matching_obj_map X (obj x) a 97 | = Hequiv ∙ (matching_obj_map X' x a) : 98 | begin 99 | unfold fn, refine nat_trans_eq, apply funext, intro y, apply funext, intro uu, 100 | cases uu, esimp at *, 101 | assert H : morphism X (hom_to y) a = 102 | natural_map ((@MO'_to_MO_map _ _ _ _ reflecting_id max_rank _ _ X) 103 | (matching_obj_map (Functor_from_C' z X) x a)) y star, unfold MO'_to_MO_map, 104 | rewrite natural_map_proj, cases y, esimp, cases x, esimp, 105 | apply H end 106 | ... = Hequiv ∙ b : ap _ p } 107 | end, 108 | apply equiv_is_fibrant Hfeq' 109 | end 110 | 111 | -- map from limit of X restricted to C' 112 | definition map_L_to_Mz (z : C) (X : C ⇒ U) [invC : invcat C] 113 | (L : cone_with_tip (Functor_from_C' z X) poly_unit) : matching_object X z := 114 | by cases L with [η, NatSq]; refine natural_transformation.mk 115 | (λ a u, η (mk (to a) (reduced_coslice_ne z a)) poly_unit.star) 116 | (λ a b f, begin apply funext, intro, cases x, refine happly (NatSq _) poly_unit.star end) 117 | 118 | 119 | -- explicit representation of the limit of X restricted to category C without z 120 | definition lim_restricted [reducible] (X : C ⇒ U) (z : C) [invC : invcat C] 121 | := Σ (c : Π y, (Functor_from_C' z X) y), 122 | Π (y y' : C_without_z z) (f : @hom (subcat_obj C _) _ y y'), 123 | ((Functor_from_C' z X) f) (c y) = c y' 124 | 125 | -- map from the limit of X restricted to C' where we use explicit representation of limit L 126 | definition map_L_to_Mz_alt (z : C) (X : C ⇒ U.{u}) [invC : invcat C] 127 | (L : lim_restricted X z) : matching_object X z := 128 | match L with 129 | | ⟨η, NatSq⟩ := 130 | by refine natural_transformation.mk 131 | (λ a u, η (mk _ (reduced_coslice_ne z a))) 132 | (λ a b f, funext (λ u, NatSq _ _ _)) 133 | end 134 | 135 | open eq.ops 136 | 137 | definition lift_to_rc [reducible] {z : C} (y : C_without_z z) (f : z ⟶ obj y): z//C := 138 | red_coslice_obs.mk (obj y) f (λ p a, prop y p⁻¹) 139 | 140 | definition singleton_contr_fiberₛ {E B : Type} {p : E → B} : (Σ b, fibreₛ p b) ≃ₛ E := 141 | begin 142 | refine equiv.mk (λ (p' : (Σ b, fibreₛ p b)), p'.2.1) (λ e, ⟨p e, ⟨e,rfl⟩⟩) _ (λx, rfl), 143 | unfold function.left_inverse, intros, cases x with [p1, p2], cases p2 with [p21,p22], 144 | esimp, induction p22 using eq.drec, congruence 145 | end 146 | 147 | definition two_piece_limit_pullback_p_q_equiv [invC : invcat C] (X : C ⇒ U.{u}) (z : C) : 148 | (Σ (c_z : X z) (c : (Π y : C_without_z z, X y)), 149 | (Π (y : C_without_z z) (f : z ⟶ obj y ), X f c_z = c y) × 150 | (Π (y y' : C_without_z z) (f : @hom (subcat_obj _ _) _ y y'), (Functor_from_C' z X) f (c y) = c y')) 151 | ≃ₛ 152 | (Σ (c_z : X z) d, !map_L_to_Mz_alt d = !matching_obj_map c_z) 153 | := 154 | begin 155 | refine @sigma_congr₂ _ _ _ _, intro c_z, 156 | refine equiv.mk _ _ _ _, 157 | { intro w, refine ⟨_,_⟩, refine ⟨w.1, prod.pr2 w.2⟩, 158 | unfold map_L_to_Mz_alt, refine nat_trans_eq, unfold natural_map, 159 | apply funext, intro y, apply funext, intro u, unfold matching_obj_map, 160 | esimp, cases w with [c,p2], cases p2 with [p_l, p_r], esimp at *, symmetry, 161 | apply (p_l (subcat_obj.mk (to y) (reduced_coslice_ne z y)) (hom_to y)) }, 162 | { intro w, 163 | cases w with [d, Heq], unfold lim_restricted at d, cases d with [c, Heq'], existsi c, 164 | split, 165 | { intros, 166 | assert H: natural_map (!map_L_to_Mz_alt ⟨c, Heq'⟩) (lift_to_rc y f) star = 167 | natural_map (!matching_obj_map c_z) (lift_to_rc y f) star, begin rewrite Heq end, 168 | assert lemma1 : natural_map (map_L_to_Mz_alt z X ⟨c,Heq'⟩) (lift_to_rc y f) star = c y, 169 | begin unfold map_L_to_Mz_alt, unfold natural_map, cases y, esimp end, 170 | have lemma2 : (natural_map (matching_obj_map X z c_z)) (lift_to_rc y f) star = X f c_z, from rfl, 171 | symmetry, 172 | apply #eq.ops lemma1⁻¹ ⬝ H ⬝ lemma2 }, 173 | { apply Heq'}}, 174 | { intro w, cases w with [c, H], cases H with [p1,p2], apply sigma_eq_congr, existsi rfl, apply apd }, 175 | { intro w, cases w with [c, H], apply sigma_eq_congr, refine ⟨_,_⟩, cases c with [p1,p2], 176 | apply sigma_eq_congr, existsi rfl, refine apd _ _, apply p2, reflexivity, apply proof_irrel } 177 | end 178 | 179 | definition lim_two_pieces_eq 180 | { X : C ⇒ U} 181 | { z : C } 182 | { c_z : X z} 183 | { c : Π y, X (obj y) } 184 | {a b : (Π (y : C_without_z z) (f : z ⟶ obj y ), X f c_z = c y) × 185 | (Π (y y' : C_without_z z) (f : @hom (subcat_obj _ _) _ y y'), (Functor_from_C' z X) f (c y) = c y')} : a = b := 186 | begin cases a, cases b, congruence end 187 | 188 | definition limit_two_piece_limit_equiv [invC : invcat C] {n' : ℕ} 189 | (ψ : C ≃ₛ fin (succ n')) 190 | {z : C} 191 | {φ : C ⇒ ℕop} 192 | (reflecting_id : id_reflect φ) 193 | (max_rank : Πy, φ y ≤ φ z) 194 | (X : C ⇒ U.{u}) : 195 | (Σ (c : Π y, X y), Π y y' f, morphism X f (c y) = c y') 196 | ≃ₛ 197 | (Σ (c_z : X z) (c : (Π y : C_without_z z, X y)), 198 | (Π (y : C_without_z z) (f : z ⟶ obj y ), X f c_z = c y) × 199 | (Π (y y' : C_without_z z) (f : @hom (subcat_obj _ _) _ y y'), (Functor_from_C' z X) f (c y) = c y')) 200 | := begin 201 | refine equiv.mk _ _ _ _, 202 | { intro, cases a with [p1, p2], refine ⟨p1 z, ⟨λ y, p1 y,(λ y' f, p2 z y' f, λ y y' f, p2 y y' f)⟩⟩ }, 203 | { intro, cases a with [c_z, p'], cases p' with [p2, p3], cases p3 with [l_z, l_y], refine ⟨_,λ y y' f, _⟩, intro y'', 204 | have Heq : decidable (y'' = z), from @fincat.has_decidable_eq _ (⟨_,ψ⟩) _ _, 205 | cases Heq with [y_eq_z, y_ne_z], 206 | { rewrite y_eq_z, assumption }, 207 | { refine p2 (mk y'' _), assumption }, 208 | intros, 209 | -- now, we have 4 cases to consider: y=z, y≠z, y'=z, y'≠z 210 | generalize (@fincat.has_decidable_eq C ⟨_,ψ⟩ y z), 211 | generalize (@fincat.has_decidable_eq C ⟨_,ψ⟩ y' z), 212 | intros deq_y deq_y', 213 | cases deq_y' with [y_eq_z, y_ne_z]: 214 | cases deq_y with [y'_eq_z, y'_ne_z], 215 | { cases y_eq_z, cases y'_eq_z, esimp, rewrite (endomorphism_is_id f), refine happly (respect_id _ _) _}, 216 | { cases y_eq_z, esimp, apply l_z}, 217 | { cases y'_eq_z, esimp, exfalso, apply (@no_incoming_non_id_arrows _ z φ max_rank _), existsi y, existsi f, assumption }, 218 | { esimp, apply l_y }, 219 | }, 220 | 221 | { unfold left_inverse, esimp, intros, cases x with [x, H], esimp, congruence, apply funext, intro y, 222 | cases @fincat.has_decidable_eq C (⟨_,ψ⟩) y z with [y_eq_z, y_ne_z], 223 | { cases y_eq_z, esimp }, {esimp }}, 224 | 225 | { unfold right_inverse, unfold left_inverse, esimp, intro y, cases y with [c_z, Hs], 226 | esimp, cases Hs with [p1, p2], esimp, cases p2, esimp, 227 | cases (@has_decidable_eq C ⟨succ n', ψ⟩ z z) with [y_eq_z, y_ne_z], esimp, congruence, 228 | apply sigma_eq_congr, 229 | refine ⟨_,lim_two_pieces_eq⟩, 230 | { intros, apply reflecting_id, assumption }, 231 | { apply funext, intro y', 232 | cases @fincat.has_decidable_eq C (⟨_,ψ⟩) y' z with [y'_eq_z, y'_ne_z], esimp, 233 | cases y' with [y'', p'], esimp at *, exfalso, apply p', apply y'_eq_z, 234 | esimp, cases y', congruence }, 235 | exfalso, apply y_ne_z, reflexivity 236 | } 237 | end 238 | 239 | definition fibration_domain_is_fibrant {E : Type} {B : Fib} (p : E → B) [isfibr_p : is_fibration p]: 240 | is_fibrant E := @equiv_is_fibrant (Σ b x, p x = b) _ singleton_contr_fiberₛ _ 241 | 242 | definition fincat_0_limit_unit_equiv [φ : C ≃ₛ fin 0 ] (X : C ⇒ U) : Nat(𝟙,X) ≃ₛ poly_unit := 243 | begin 244 | cases φ with [f,g,l,r], 245 | refine equiv.mk (λ x, star) _ _ _, 246 | { intros, esimp, refine mk _ _, intros x, exfalso, apply (false_of_fin_zero (f x)), 247 | intros a b u, exfalso, apply (false_of_fin_zero (f a))}, 248 | { unfold left_inverse, intro L, cases L, congruence, apply funext, intro x, exfalso, apply (false_of_fin_zero (f x))}, 249 | { unfold right_inverse, unfold left_inverse, intro x, cases x, reflexivity } 250 | end 251 | -------------------------------------------------------------------------------- /2ltt/finite.lean: -------------------------------------------------------------------------------- 1 | import .fibrant data.fin data.equiv .facts algebra.category .facts 2 | 3 | open nat equiv fin eq.ops sum unit prod.ops function 4 | 5 | -- facts about a type family indexed by the strict finite type 6 | 7 | definition pi_fin0_unit_equiv {X : fin 0 → Type} : (Π i, X i) ≃ unit := 8 | equiv.mk (λ x, unit.star) (λ x i, fin.elim0 i) 9 | begin unfold left_inverse, intro, apply funext, intro, apply (elim0 x_1) end 10 | begin 11 | intro, cases x, reflexivity 12 | end 13 | 14 | lemma pi_sum_fin_unit_equiv {n} {X : (unit + fin n) → Type} : (Π i, X i) ≃ (X (inl ⋆) × Π i, X (inr i)) := 15 | equiv.mk 16 | (λ x, (x (inl ⋆), λ y, x (inr y))) 17 | (λ p, λ z, sum.cases_on z (λ x1, unit.cases_on x1 p.1) (λ x2, p.2 x2)) 18 | begin 19 | unfold left_inverse, esimp, intro, apply funext, intro i, 20 | cases i, 21 | { esimp, cases a, esimp }, 22 | { esimp } 23 | end 24 | begin 25 | unfold right_inverse, unfold left_inverse, apply prod.eta 26 | end 27 | 28 | definition pi_sum_fin_unit_equiv' {n} (Heq : fin n + unit ≃ fin (nat.succ n)) 29 | {X : fin (nat.succ n) → Type} 30 | : (Π i, X (Heq ∙ i)) ≃ (X (Heq ∙ (inr ⋆)) × Π i, X (Heq ∙ (inl i))) := 31 | equiv.mk 32 | (λ x, (x (inr ⋆), λ y, x (inl y))) 33 | (λ p, λ z, sum.cases_on z (λ x2, p.2 x2) (λ x1, unit.cases_on x1 p.1)) 34 | begin 35 | unfold left_inverse, esimp, intro x, apply funext, intro i, 36 | cases i, esimp, cases a, esimp 37 | end 38 | begin 39 | unfold right_inverse, unfold left_inverse, apply prod.eta 40 | end 41 | 42 | -- ref:lem:finite-cofibrant 43 | -- Lemma 3.23 (iii) 44 | definition finite_cofibrant {n : ℕ} {X : fin n → Fib} 45 | : is_fibrant (Π i, X i) := 46 | begin 47 | induction n with n IHn, 48 | { apply (equiv_is_fibrant (equiv.symm pi_fin0_unit_equiv)) }, 49 | { have HeqFinSum : fin n + unit ≃ fin (succ n), from (fin_sum_unit_equiv n), 50 | apply equiv_is_fibrant, 51 | apply pi_congr₁, 52 | apply equiv_is_fibrant, apply (equiv.symm (pi_sum_fin_unit_equiv' HeqFinSum))} 53 | end 54 | 55 | namespace fin 56 | 57 | definition lift_down {n : ℕ} (i : fin (nat.succ n)) (Hne : i ≠ maxi) : fin n := fin.mk (val i) (lt_max_of_ne_max Hne) 58 | 59 | definition lift_succ_lift_down_inverse {n : ℕ} {i : fin (nat.succ n)} {Hne : i ≠ maxi} : 60 | (lift_succ (lift_down i Hne)) = i := 61 | begin cases i, esimp end 62 | 63 | open decidable equiv.ops 64 | 65 | -- inspired by Paolo Capriotti's Agda implementation 66 | 67 | definition fin_transpose {n} (i j k : fin n) : fin n := 68 | match fin.has_decidable_eq _ i k with 69 | | inl _ := j 70 | | inr _ := match fin.has_decidable_eq _ j k with 71 | | inl _ := i 72 | | inr _ := k 73 | end 74 | end 75 | 76 | lemma fin_transpose_invol {n} {i j k : fin n} : fin_transpose i j (fin_transpose i j k) = k := 77 | begin 78 | unfold fin_transpose, 79 | cases (fin.has_decidable_eq n i k) with [i_eq_k, i_ne_k]: esimp, 80 | cases (fin.has_decidable_eq n i j) with [i_eq_j, j_ne_k]: esimp, apply i_eq_j⁻¹ ⬝ i_eq_k, 81 | cases (fin.has_decidable_eq n j j) with [j_eq_j, j_ne_j]: esimp, assumption, apply absurd rfl j_ne_j, 82 | cases (fin.has_decidable_eq n j k) with [j_eq_k, j_ne_k]: esimp, 83 | cases (fin.has_decidable_eq n i i) with [i_eq_i, i_ne_i]: esimp, assumption, 84 | cases (fin.has_decidable_eq n j i) with [j_eq_i, j_ne_i]: esimp: apply absurd rfl i_ne_i, 85 | cases (fin.has_decidable_eq n i k) with [i_eq_k, i_ne_k]: esimp, rewrite i_eq_k at i_ne_k, apply absurd rfl i_ne_k, 86 | cases (fin.has_decidable_eq n j k) with [j_eq_k, j_ne_k]: esimp, apply absurd j_eq_k j_ne_k 87 | end 88 | 89 | definition fin_transpose_equiv [instance] {n} (i j : fin n) : fin n ≃ fin n:= 90 | equiv.mk (fin_transpose i j) (fin_transpose i j) (λ k, fin_transpose_invol) (λ k, fin_transpose_invol) 91 | 92 | definition fin_transpose_β₁ {n : ℕ} {i j : fin n} : fin_transpose i i j = j := 93 | begin 94 | unfold fin_transpose, 95 | cases (fin.has_decidable_eq n i j) with [i_eq_j, i_ne_j]: esimp, apply i_eq_j 96 | end 97 | 98 | definition fin_transpose_β₂ {n : ℕ} {i j : fin n} : fin_transpose i j j = i := 99 | begin 100 | unfold fin_transpose, 101 | cases (fin.has_decidable_eq n i j) with [i_eq_j, i_ne_j]: esimp, apply i_eq_j⁻¹, 102 | cases (fin.has_decidable_eq n j j) with [j_eq_j, j_ne_j]: esimp, apply absurd rfl j_ne_j 103 | end 104 | 105 | definition fin_transpose_inj {n : ℕ} {i j : fin n} : injective (fin_transpose i j) := 106 | begin 107 | refine (injective_of_left_inverse (@left_inv (fin n) (fin n) (fin_transpose_equiv _ _))) 108 | end 109 | 110 | -- removing any element from fin is equivalent to removing maximal element 111 | -- we use transpositions defined above to prove this 112 | definition fin_remove_max_equiv [instance] {n : ℕ} (z : fin (nat.succ n)) : 113 | (Σ i : fin (nat.succ n), i ≠ z) ≃ (Σ i : fin (nat.succ n), i ≠ maxi) := 114 | begin 115 | assert H : Π (x : fin (nat.succ n)), x = z ≃ (fin_transpose maxi z x = maxi), 116 | begin 117 | intros, refine equiv.mk (λ Heq, Heq⁻¹ ▹ fin_transpose_β₂) _ _ _, 118 | { intros, assert Heq : fin_transpose maxi z x = fin_transpose maxi z z, begin rewrite fin_transpose_β₂, assumption end, 119 | apply fin_transpose_inj, assumption }, 120 | { unfold left_inverse, intro, esimp }, 121 | { unfold right_inverse, intro, esimp } 122 | end, 123 | refine @sigma_congr _ _ _ _ (fin_transpose_equiv maxi z) (λ x, @pi_congr _ _ _ _ (H _) (λ y, equiv.refl _)) 124 | end 125 | 126 | open sigma.ops 127 | 128 | -- removing maximal element gives us a set with smaller cardinality 129 | -- we use it as a base case to remove any element and get a set with smaller cardinality 130 | definition fin_remove_max [instance] {n : ℕ} : (Σ i : fin (nat.succ n), i ≠ maxi) ≃ fin n := 131 | begin 132 | refine equiv.mk (λ j, lift_down _ j.2) (λi,⟨lift_succ i,lift_succ_ne_max⟩) _ _, 133 | { unfold left_inverse, intro j, cases j, congruence, 134 | rewrite lift_succ_lift_down_inverse }, 135 | { unfold right_inverse, unfold left_inverse, intro i, cases i, unfold lift_down } 136 | end 137 | 138 | -- now we can remove any element from fin using equivalences defined above 139 | definition fin_remove_equiv {n : ℕ } (z : fin (nat.succ n)) 140 | : (Σ i : fin (nat.succ n), i ≠ z) ≃ fin n := fin_remove_max ∘ (fin_remove_max_equiv z) 141 | end fin 142 | 143 | -- some facts about finite categories 144 | 145 | namespace fincat 146 | universe variables u v 147 | variables {C C' : Category.{u v}} 148 | open category sigma.ops eq equiv.ops 149 | 150 | definition is_obj_finite [class] (C : Category) := Σ n, C ≃ fin n 151 | definition is_obj_finite_eq [finC : is_obj_finite C] := finC.2 152 | 153 | attribute is_obj_finite_eq [instance] 154 | 155 | definition fincat_ne_maxi {n : ℕ} {z : C} {f : C → fin (succ n)} (inj_f : injective f) 156 | (max_z : f z = maxi) {o : C} (p : o ≠ z) : f o ≠ (maxi : fin (succ n)) := 157 | begin 158 | unfold ne, unfold not, intro, 159 | apply p, rewrite -max_z at a, apply inj_f a 160 | end 161 | 162 | definition eq_of_finN_eq [finC : is_obj_finite C] {c c' : C} : c = c' → to_fun (fin (finC.1)) c = to_fun (fin (finC.1)) c' := 163 | begin 164 | intro Heq, cases finC with [n, φ], esimp, apply (ap _ Heq) 165 | end 166 | 167 | definition finN_eq_of_eq [finC : is_obj_finite C] {c c' : C} : to_fun (fin (finC.1)) c = to_fun (fin (finC.1)) c' → c = c' := 168 | begin 169 | intro Heq, cases finC with [n, φ], cases φ with [f,g,l,r], esimp at *, 170 | apply (injective_of_left_inverse l), apply Heq 171 | end 172 | 173 | definition eq_iff_finN_eq [finC : is_obj_finite C] {c c' : C} : 174 | to_fun (fin (finC.1)) c = to_fun (fin (finC.1)) c' ↔ c = c' := iff.intro finN_eq_of_eq eq_of_finN_eq 175 | 176 | definition has_decidable_eq [instance] [finC : is_obj_finite C] {c c' : C} : decidable (c = c') := 177 | decidable_of_decidable_of_iff (fin.has_decidable_eq _ _ _) eq_iff_finN_eq 178 | 179 | definition fincat_ob_remove_fin_equiv {n : ℕ} (z : C) [φ : C ≃ fin (succ n)] : (Σ c, c ≠ z) ≃ fin n := 180 | begin 181 | cases φ with [f,g,l,r], esimp at *, 182 | assert Hequiv: (Σ c, c ≠ z) ≃ Σ (i : fin (succ n)), i ≠ f z, 183 | begin refine equiv.mk _ _ _ _, 184 | intro j, cases j with [c, p_ne], existsi f c, intros, apply p_ne, apply (injective_of_left_inverse l), assumption, 185 | intro i, cases i with [i', p_ne], existsi g i', intros, apply p_ne, rewrite -a, apply (r i')⁻¹, 186 | unfold left_inverse, intro x, cases x with [c, p_ne], esimp, congruence, apply l, 187 | unfold right_inverse, unfold left_inverse, unfold injective_of_left_inverse, intro x, 188 | cases x with [i, p_ne], esimp, congruence, apply r end, 189 | apply (fin.fin_remove_equiv _) ∘ Hequiv 190 | end 191 | end fincat 192 | -------------------------------------------------------------------------------- /2ltt/inverse.lean: -------------------------------------------------------------------------------- 1 | -- inverse categories 2 | 3 | import algebra.category data.nat 4 | import algebra.category.constructions 5 | 6 | import .fibrant .finite 7 | 8 | import open nat 9 | 10 | open nat.le functor 11 | open category category.ops eq.ops 12 | 13 | definition nat_cat_op [instance] : category ℕ := 14 | ⦃ category, 15 | hom := λ a b, a ≥ b, 16 | comp := λ a b c, @nat.le_trans c b a, 17 | ID := nat.le_refl, 18 | assoc := λ a b c d h g f, eq.refl _, 19 | id_left := λ a b f, eq.refl _, 20 | id_right := λ a b f, eq.refl _ ⦄ 21 | 22 | definition ℕop := Mk nat_cat_op 23 | 24 | lemma hom_ℕop_id {C : Category} {x : ℕop} {f : x ⟶ x} : f = id := rfl 25 | 26 | -- ref:def:inverse-category 27 | -- Definition 4.1 28 | namespace invcat 29 | open sigma.ops iff 30 | definition id_reflect {C D: Category} (φ : C ⇒ D) := 31 | Π ⦃x y : C⦄ (f : x ⟶ y), (Σ (q : φ x = φ y), q ▹ φ f = id) → Σ (p : x = y), p ▹ f = id 32 | 33 | definition id_reflect' {C D: Category} := 34 | Π (φ : C ⇒ D) ⦃x y : C⦄ (f : x ⟶ y), (Σ (q : φ x = φ y), q ▹ φ f = id) → Σ (p : x = y), p ▹ f = id 35 | 36 | definition transport {A : Type} := @eq.rec_on A 37 | notation x ≡ y := eq x y 38 | 39 | definition id_reflect'' {C D: Category} := 40 | Π (φ : C ⇒ D) {x y : C} (f : hom x y), 41 | (Σ (q : φ y ≡ φ x), transport q (φ f) ≡ id) → Σ (p : y ≡ x), 42 | @transport _ _ _ _ p f ≡ id 43 | 44 | 45 | -- definition if "refect identity" property specific to ℕop. 46 | -- Though this definition doesn't require for φ(f) to be an identity, we will show that 47 | -- id_reflect(C,ℕop) and id_reflect_ℕop(C) are logically equivalent 48 | definition id_reflect_ℕop {C : Category} (φ : C ⇒ ℕop) := 49 | Π ⦃x y : C⦄ (f : x ⟶ y), φ x = φ y → (Σ (p : x = y), p ▹ f = id) 50 | 51 | definition id_reflect_ℕop' {C : Category} (φ : C ⇒ ℕop) := 52 | Π {x y : C} (f : x ⟶ y), 53 | φ x ≡ φ y → (Σ (p : x ≡ y), transport p f ≡ id) 54 | 55 | -- have to pack functor with the property that it reflects identities, 56 | -- because functor itself is not a type class 57 | structure has_idreflect [class] (C D : Category) := 58 | (φ : C ⇒ D) 59 | (reflecting_id : id_reflect φ) 60 | 61 | section id_reflect_ℕop_iff 62 | 63 | -- showing that id_reflect_ℕop φ ↔ id_reflect φ 64 | -- NOTE: (Danil) I couldn't find instance of iff for Type 65 | definition map_id_reflect_ℕop {C : Category} (φ : C ⇒ ℕop): 66 | id_reflect_ℕop φ → id_reflect φ := 67 | begin 68 | intros id_r_ℕop x y f s, cases s with [q, Heq], cases id_r_ℕop f q with [p₁, p₂], 69 | existsi p₁, cases p₁, esimp at *, apply p₂ 70 | end 71 | 72 | definition map_ℕop_id_reflect {C : Category} (φ : C ⇒ ℕop): 73 | id_reflect φ → id_reflect_ℕop φ := 74 | begin intros id_r x y f q, 75 | -- here we use the fact that any morphism x ⟶ x in ℕop only can be an identity morphism 76 | have f_is_id : q ▹ morphism φ f = id, from rfl, 77 | cases id_r f ⟨q,f_is_id⟩ with [p₁, p₂], existsi p₁, cases p₁, esimp at *, apply p₂ 78 | end 79 | 80 | end id_reflect_ℕop_iff 81 | 82 | structure invcat [class] (C : Category) := 83 | (reflecting_id_ℕop : has_idreflect C ℕop) 84 | 85 | end invcat 86 | 87 | -- some facts about inverse categories 88 | section invcat_facts 89 | open invcat function sigma.ops 90 | 91 | definition endomorphism_is_id {C : Category} [invC : invcat C] {c : C} (f : c ⟶ c) : f = id := 92 | begin cases invC with H, cases H with [φ, id_r], apply (id_r f ⟨rfl,rfl⟩).2 end 93 | 94 | lemma idreflect_inj_hom {C : Category} {x y : C} [idr : has_idreflect C ℕop] (f : x ⟶ y) : 95 | (has_idreflect.φ C ℕop) x = (has_idreflect.φ C ℕop) y → x = y := 96 | begin cases idr with [φ, idr_φ], esimp, intro H, cases (idr_φ f ⟨H,rfl⟩), assumption end 97 | 98 | definition has_le_ℕop [instance] : has_le ℕop := nat_has_le 99 | definition has_lt_ℕop [instance] : has_lt ℕop := nat_has_lt 100 | 101 | definition strict_order_ℕop [instance] : strict_order ℕop := 102 | strict_order.mk nat.lt begin intros, refine (@lt.irrefl ℕ _ a a_1) end 103 | begin intros a b c, apply nat.lt_trans end 104 | 105 | definition weak_order_ℕop [instance] : weak_order ℕop := 106 | weak_order.mk nat.le (@le.refl ℕ _) (λ a b c, @nat.le_trans a b c) (λ a b, nat.le_antisymm) 107 | 108 | definition no_incoming_non_id_arrows {C : Category.{1 1}} 109 | (z : C) {φ : C ⇒ ℕop} {max_rank : ∀ y, φ y ≤ φ z} {reflecting_id : id_reflect φ} 110 | : ¬ ∃ y : C, ∃ (f : y ⟶ z), y ≠ z := 111 | begin intro H, cases H with [y, s], cases s with [f, y_ne_z], unfold id_reflect at *, 112 | have H : φ y ≥ φ z, from φ f, 113 | cases reflecting_id f ⟨le.antisymm (max_rank y) H, rfl⟩ with [p, Heq], 114 | apply y_ne_z, assumption 115 | end 116 | end invcat_facts 117 | 118 | open invcat 119 | open unit 120 | 121 | definition triv_funct : Category_one ⇒ ℕop := 122 | functor.mk (λ (x : unit), zero) (λ a b p, id) (λa, eq.refl _) (λa b c p q, eq.refl _) 123 | 124 | definition triv_funct_id_reflect [instance] : has_idreflect Category_one ℕop := 125 | has_idreflect.mk 126 | triv_funct 127 | begin 128 | intros x y f, 129 | cases x, cases y, intro p, 130 | existsi (eq.refl _), 131 | cases f, reflexivity 132 | end 133 | 134 | definition triv_cat_inverse [instance] : invcat Category_one := invcat.mk _ 135 | 136 | structure subcat_obj (C : Category) (p : objects C → Prop) := 137 | (obj : objects C) 138 | (prop : p obj) 139 | 140 | open subcat_obj function 141 | attribute subcat_obj.obj [coercion] 142 | 143 | 144 | section subcategory 145 | variables {C : Category.{1 1}} 146 | {D : Category} 147 | 148 | definition subcat [instance] (C : Category) (p : C → Prop) : category (subcat_obj C p) := 149 | ⦃ category, 150 | hom := λ (a b : subcat_obj C p), obj a ⟶ obj b, 151 | comp := λ a b c, @comp _ _ _ _ _, 152 | ID := λ (a : subcat_obj C p), ID (obj a), 153 | assoc := λ a b c d h g f, assoc h g f, 154 | id_left := λ a b f, id_left f, 155 | id_right := λ a b f, id_right f ⦄ 156 | 157 | definition Subcat (C : Category) (p : C → Prop) : Category := Mk (subcat C p) 158 | 159 | 160 | -- A subcategory of [C] with one object removed. 161 | -- We also refer to this category as C' 162 | definition C_without_z {C : Category.{1 1}}(z : C) : Category := Mk (subcat C (λ c, c ≠ z)) 163 | 164 | -- use apply tactic, as it allows to infer correct implicits 165 | definition Functor_from_C' [reducible] (z : C) (X : C ⇒ D) : C_without_z z ⇒ D := 166 | ⦃ functor, 167 | object := λ ob, X (obj ob), 168 | morphism := λ a b f, by apply X f, 169 | respect_id := λ a, by apply respect_id X (obj a), 170 | respect_comp := λ a b c g f, by apply @respect_comp _ _ X (obj a) (obj b) (obj c) _ _ ⦄ 171 | 172 | open eq.ops function 173 | definition C_without_z_invcat [instance] (z : C) [invC : invcat C] : invcat (C_without_z z) := 174 | begin 175 | unfold C_without_z, cases invC, cases reflecting_id_ℕop, 176 | refine invcat.mk (has_idreflect.mk _ _), apply Functor_from_C' z φ, intros x y f p, 177 | cases reflecting_id f p with [p1, p2], 178 | refine ⟨_,_⟩, 179 | { cases x, cases y, congruence, apply p1 }, 180 | { cases x, cases y, esimp at *, induction p1 using eq.drec, esimp at *, apply p2} 181 | end 182 | 183 | open equiv equiv.ops sigma sigma.ops 184 | 185 | definition C_without_z_sigma_equiv {C : Category} (z : C) : C_without_z z ≃ Σ (c : C), c ≠ z := 186 | equiv.mk (λ c', ⟨obj c', prop c'⟩) (λc, mk c.1 c.2) begin intros, cases x, esimp, end begin intros, cases x, esimp end 187 | 188 | definition C_without_z_is_obj_finite [instance] {n : ℕ} (z : C) [φ : objects C ≃ fin (nat.succ n)] 189 | : objects (C_without_z z) ≃ fin n := (fincat.fincat_ob_remove_fin_equiv z) ∘ (C_without_z_sigma_equiv z) 190 | 191 | end subcategory 192 | -------------------------------------------------------------------------------- /2ltt/limit.lean: -------------------------------------------------------------------------------- 1 | import data.equiv algebra.category .facts 2 | 3 | open natural_transformation sigma.ops 4 | open category function equiv 5 | 6 | notation X `≃ₛ` Y := equiv X Y 7 | 8 | -- Given categories J and C, we have a canonical functor [const_funct] from the category C to the functor category C^J. 9 | -- Here, we do not define this functor in full, but we define the object and morphism part. 10 | -- These are called [const_funct_obj] and [const_funct_morph]. 11 | -- It would be very easy to complete this to an actual functor, but we do not need this. 12 | 13 | definition const_funct_obj [reducible] [unfold_full] (J C : Category) (c : C) : J ⇒ C := 14 | ⦃ functor, 15 | object := λ i, c, 16 | morphism := λ i j g, id, 17 | respect_id := λ i, eq.refl _, 18 | respect_comp := λi j k f g, by rewrite id_left ⦄ 19 | 20 | definition const_funct_morph [reducible] (J C : Category) (c d : C) (f : c ⟶ d) : (const_funct_obj J C c) ⟹ (const_funct_obj J C d) 21 | := mk (λ j, f) 22 | begin intros, esimp, rewrite id_left, rewrite id_right end 23 | 24 | 25 | -- Given categories J and C as before, and a functor D : J ⇒ C, we have a category of cones. 26 | -- We present it as "fibred" over C. For each object [tip : C] we have a category (here only a type) of cones 27 | -- with "tip" [tip]. This is just the type of natural transformations between the functor constantly [tip] and D. 28 | definition cone_with_tip [reducible] [unfold_full] {J C : Category} (D : J ⇒ C) (tip : C) := const_funct_obj _ _ tip ⟹ D 29 | 30 | open functor 31 | -- For [f : tip₁ ⟶ tip₂], we have a function between the type of cones with tip tip₂, and the onew with tip tip₁. 32 | definition cone_with_tip_functorial [reducible] [unfold_full] {J C : Category} (D : J ⇒ C) (tip₁ tip₂ : C) (f : tip₁ ⟶ tip₂) (c₂ : cone_with_tip D tip₂) 33 | : cone_with_tip D tip₁ 34 | := natural_transformation.compose c₂ (const_funct_morph J C tip₁ tip₂ f) 35 | 36 | -- A cone is a tip together with a cone under this tip. 37 | definition cone [reducible] [unfold_full] {J C : Category} (D : J ⇒ C) := Σ c, cone_with_tip D c 38 | 39 | -- morphisms of cones. 40 | structure cone_hom {J C : Category} {D : J ⇒ C} (c : cone D ) (c' : cone D) : Type := 41 | (chom : c.1 ⟶ c'.1) 42 | (commute_triangle : ∀ i, natural_map c.2 i = natural_map c'.2 i ∘ chom) 43 | 44 | open cone_hom 45 | 46 | 47 | 48 | definition cone_hom_eq {J C : Category} {D : J ⇒ C } {c c' : cone D} 49 | {f f': cone_hom c c'} (p : chom f = chom f') : f = f' := 50 | begin cases f, cases f', cases p, reflexivity end 51 | 52 | 53 | -- Lemma stating that equality between nat. transf. is equality of morphisms (laws trivially equal) 54 | -- Danil : It seems that the second goal (laws) is resolved automatically. So, we could use congruence 55 | -- tactic directly to compare two natural transformations, but I'll leave this here as an example. 56 | definition cone_with_tip_eq {J C : Category} (D : J ⇒ C) (tip : C) (c₁ c₂ : cone_with_tip D tip) : 57 | (natural_map c₁ = natural_map c₂) → c₁ = c₂ 58 | := begin intros NatMapEq, cases c₁, cases c₂, congruence, esimp at *, apply NatMapEq end 59 | 60 | 61 | definition cone_hom_comp {J C : Category} {D : J ⇒ C } {c c' c'': cone D} 62 | (f' : cone_hom c' c'') (f : cone_hom c c') := 63 | ⦃ cone_hom, chom := chom f' ∘ chom f, 64 | commute_triangle := λ i, by rewrite [assoc, -commute_triangle f', commute_triangle f] ⦄ 65 | 66 | definition cone_category [instance] [reducible] [unfold_full] {J C : Category} (D : J ⇒ C) : category (cone D) := 67 | ⦃ category, 68 | hom := λ c c', cone_hom c c', 69 | comp := λ a b c f g, cone_hom_comp f g, 70 | ID := λi, ⦃ cone_hom, chom := id, commute_triangle := λ i, by rewrite id_right ⦄, 71 | assoc := λ a b c p q h g, by cases h; cases q; cases g; apply cone_hom_eq; apply assoc, 72 | id_left := λ a b f, cone_hom_eq (id_left (chom f)), 73 | id_right := λ a b f, cone_hom_eq (id_right (chom f)) ⦄ 74 | 75 | definition ConeCat [reducible] {J C : Category} (D : J ⇒ C) : Category := category.Mk (cone_category D) 76 | 77 | set_option formatter.hide_full_terms false 78 | 79 | 80 | -- CAVEAT: I (Nicolai) have changed this. 81 | structure is_terminal [class] {C : Category} (c : C) := 82 | (term_hom : ∀ c', hom c' c) 83 | (unique_term_hom : ∀ c' (f : hom c' c), f = term_hom c') 84 | 85 | structure has_terminal_obj [class] (C : Category) := 86 | (terminal : C) 87 | (is_terminal_obj : is_terminal terminal) 88 | 89 | open has_terminal_obj 90 | 91 | definition limit {J C : Category} (D : J ⇒ C) : Type := has_terminal_obj (ConeCat D) 92 | 93 | definition limit_obj [reducible] [unfold_full] {J C : Category} {D : J ⇒ C} : limit D → objects C 94 | | limit_obj (has_terminal_obj.mk c _) := c.1 95 | 96 | notation `Nat` `(` F `,` G `)` := F ⟹ G 97 | 98 | open eq.ops 99 | 100 | open functor natural_transformation 101 | 102 | definition nat_transf_sigma_iso {C D : Category} {F G : C ⇒ D} : 103 | F ⟹ G ≃ₛ Σ (η : Π(a : C), hom (F a) (G a)), (Π{a b : C} (f : hom a b), G f ∘ η a = η b ∘ F f) := 104 | equiv.mk (λ N, match N with 105 | |mk η NatSq := ⟨η, NatSq⟩ 106 | end) 107 | (λ S, match S with 108 | | ⟨η, NatSq⟩ := mk η NatSq 109 | end) 110 | begin unfold function.left_inverse, intro x, cases x, esimp end 111 | begin unfold function.right_inverse, unfold function.left_inverse, intro x, cases x, esimp end 112 | 113 | 114 | open eq 115 | set_option pp.all true 116 | 117 | open equiv poly_unit 118 | 119 | definition poly_unit_arrow_equiv [instance] [simp] (A : Type) : (poly_unit → A) ≃ A := 120 | mk (λ f, f star) (λ a, (λ u, a)) 121 | (λ f, funext (λ x, by cases x; reflexivity)) 122 | (λ u, rfl) 123 | 124 | definition to_unit [reducible] [unfold_full] {C : Category} {X : C ⇒ Type_category } 125 | (f : Π a, poly_unit → X a) y := f y star 126 | 127 | definition pi_unit_arrow_equiv {C : Category} {X : C ⇒ Type_category } : 128 | (Π a, object (const_funct_obj C Type_category poly_unit) a⟶ X a) ≃ Π y, X y := 129 | begin 130 | esimp, refine equiv.mk to_unit (λ f y x, f y) _ (λx, rfl), 131 | unfold function.left_inverse, intros, apply funext, intros, apply funext, intros, cases x_2, reflexivity 132 | end 133 | 134 | definition nat_unit_sigma_equiv [instance] {C : Category.{1 1}} {X : C ⇒ Type_category} : 135 | (const_funct_obj C Type_category poly_unit ⟹ X) ≃ₛ 136 | Σ (c : Π y, X y), Π (y y' : C) (f : y ⟶ y'), (X f) (c y) = c y' := 137 | begin 138 | apply equiv.trans (nat_transf_sigma_iso), 139 | apply @sigma_congr, 140 | 141 | intros ff, 142 | apply @pi_congr₂, intro, apply @pi_congr₂, intro b, apply @pi_congr₂, intro f', 143 | esimp at *, rewrite id_right, 144 | refine equiv.mk _ _ _ _, 145 | apply pi_unit_arrow_equiv, 146 | { intro p, apply happly p }, 147 | 148 | { intro p, esimp at *, 149 | apply funext, intro x, cases x, apply p }, 150 | 151 | { unfold function.left_inverse, intro, esimp }, 152 | 153 | { unfold function.right_inverse, unfold function.left_inverse, intro p, esimp }, 154 | end 155 | 156 | definition one_funct [reducible] [unfold_full] {C : Category.{1}} := const_funct_obj C Type_category poly_unit 157 | notation `𝟙` := one_funct 158 | 159 | -- binary product as a limit 160 | 161 | open bool prod eq.ops 162 | 163 | -- it's a bit akward to use definition of Category_two from library, 164 | -- so we use this definition 165 | definition cat_two : category bool := 166 | ⦃ category, 167 | hom := λ a b, a = b, 168 | comp := λ a b c p q, q ⬝ p, 169 | ID := λ a, eq.refl _, 170 | assoc := λ a b c d h g f, by cases h; cases g; cases f; reflexivity, 171 | id_left := λ a b f, by cases f; reflexivity, 172 | id_right := λ a b f, by cases f; reflexivity ⦄ 173 | 174 | definition CatTwo := Mk cat_two 175 | 176 | definition c2_functor (C : Category) (A B : C) : CatTwo ⇒ C := 177 | ⦃ functor, 178 | object := bool.rec A B, 179 | morphism := (bool.rec (bool.rec (λf, id) (by contradiction)) 180 | (bool.rec (by contradiction) (λf, id))), 181 | respect_id := bool.rec rfl rfl, 182 | respect_comp := λi j k f g, by cases f: cases g: cases k: esimp: rewrite id_left ⦄ 183 | 184 | definition product {C : Category} (A B : C) (L : limit (c2_functor _ A B)) := limit_obj L 185 | 186 | -- Type_category limits 187 | 188 | --set_option pp.universes true 189 | 190 | open natural_transformation 191 | open - [notation] category 192 | 193 | open functor poly_unit 194 | 195 | universe variable u 196 | 197 | definition cone_in_pretype [reducible] {J : Category.{1}} (D : J ⇒ Type_category.{max 1 u}) : cone D := 198 | ⟨ cone_with_tip _ poly_unit, -- (const_funct_obj _ _ unit) ⟹ D , 199 | natural_transformation.mk 200 | (λ a L, natural_map L a star) 201 | (λ a b f, funext (λ L, happly (naturality L f) _)) ⟩ 202 | 203 | open function 204 | 205 | -- just for rewriting in limit_in_pretype, because projections are left not simplified sometimes 206 | definition natural_map_proj {C D : Category} (F G: functor C D) (η :Π a, F a ⟶ G a) 207 | (nat : Π⦃a b : C⦄ (f : a ⟶ b), G f ∘ η a = η b ∘ F f) : natural_map (natural_transformation.mk η nat) = η := rfl 208 | 209 | definition nat_trans_eq {C D : Category} {F G : C ⇒ D} {N M: F ⟹ G} 210 | {p : natural_map N = natural_map M} : N = M := 211 | begin cases N with [η, NatSq], cases M with [η', NatSq'], unfold natural_map at *, cases p, congruence end 212 | 213 | definition natural_map_eq {C D : Category} {F G : C ⇒ D} {N M: F ⟹ G} (p : N = M) : natural_map N = natural_map M 214 | := begin cases N with [η, NatSq], cases M with [η', NatSq'], unfold natural_map, injection p, assumption end 215 | 216 | 217 | definition limit_in_pretype {J : Category.{1}} (D : J ⇒ Type_category) : limit D := 218 | ⦃ has_terminal_obj _, 219 | terminal := cone_in_pretype D, 220 | is_terminal_obj := 221 | ⦃ is_terminal _, 222 | term_hom := λ C, mk (λ x, cone_with_tip_functorial D poly_unit C.1 (λ tt, x) C.2) (λ x, rfl), 223 | unique_term_hom := 224 | begin 225 | intros C f, 226 | cases f with [chom', ct], 227 | apply cone_hom_eq, 228 | apply funext, intro x, esimp, 229 | -- now: need to show equality of two cones with tip unit: 230 | apply cone_with_tip_eq, 231 | apply funext, intro j, 232 | apply funext, intro t, 233 | -- now: have do show an equality in D(j). 234 | esimp at *, unfold natural_transformation.compose, repeat rewrite natural_map_proj, 235 | unfold cone_in_pretype at *, 236 | rewrite natural_map_proj at ct, rewrite ct, 237 | -- we have to assert the same goal, but explicitly say that composition is just composition of functions, 238 | -- because some definitions related to class instances are not unfolded 239 | assert HH : natural_map (chom' x) j t = #function (((λ L, natural_map L j star)∘chom')∘λ tt, x) t, 240 | begin esimp, cases t, reflexivity end, 241 | exact HH 242 | end 243 | ⦄ 244 | ⦄ 245 | 246 | 247 | -- construction of the product in Pretype (aka Type_category) using the 248 | -- limit of the diagramm 𝟚 ⇒ Type_category 249 | definition product_in_pretype (A B : Type) := product A B (@limit_in_pretype CatTwo (c2_functor Type_category A B)) 250 | 251 | open prod.ops 252 | 253 | -- showing that constructed product is isomophic to the product type 254 | definition lim_prod_iso {A B : Type_category} : (product_in_pretype A B) ≃ (A × B) := 255 | begin 256 | refine equiv.mk _ _ _ _, 257 | { intro, cases a with [η, comm_tr], esimp, refine (η ff poly_unit.star, η tt poly_unit.star) }, 258 | { intro p, refine natural_transformation.mk _ _, 259 | { intro a uu, cases a, apply p.1, apply p.2 }, 260 | { intros, cases f, cases b, esimp, apply funext, intro x, reflexivity }}, 261 | { intros, refine nat_trans_eq, cases x with [η, comm_tr], 262 | rewrite natural_map_proj, rewrite natural_map_proj, 263 | apply funext, intros b, apply funext, intro uu, cases uu, cases b: esimp }, 264 | { intros, esimp, cases x, esimp } 265 | end 266 | -------------------------------------------------------------------------------- /2ltt/matching.lean: -------------------------------------------------------------------------------- 1 | import algebra.category 2 | import .inverse .limit .fibrant .finite 3 | 4 | open sigma category eq.ops function 5 | 6 | -- Lean's types are pretypes in our formalisation, so the category of types [Types_category] from the standard library is actually the category of pretypes. 7 | definition U := category.Type_category 8 | 9 | namespace reduced_coslice 10 | structure coslice_obs {ob : Type} (C : category ob) (a : ob) := 11 | (to : ob) 12 | (hom_to : hom a to) 13 | 14 | open coslice_obs 15 | 16 | structure red_coslice_obs {A : Type} (C : category A) (c : A) extends coslice_obs C c := 17 | (rc_non_id_hom : Π (p : c = to), not (p ▹ hom_to = category.id)) 18 | 19 | 20 | definition reduced_coslice {ob : Type} (C : category ob) (c : ob) : category (red_coslice_obs C c) := 21 | ⦃ category, 22 | hom := λa b, Σ(g : hom (to a) (to b)), g ∘ hom_to a = hom_to b, 23 | comp := λ a b c g f, 24 | ⟨ (pr1 g ∘ pr1 f), 25 | (show (pr1 g ∘ pr1 f) ∘ hom_to a = hom_to c, 26 | proof 27 | calc 28 | (pr1 g ∘ pr1 f) ∘ hom_to a = pr1 g ∘ (pr1 f ∘ hom_to a): eq.symm !assoc 29 | ... = pr1 g ∘ hom_to b : {pr2 f} 30 | ... = hom_to c : {pr2 g} 31 | qed) ⟩, 32 | ID := (λ a, ⟨ id, !id_left ⟩), 33 | assoc := (λ a b c d h g f, dpair_eq !assoc !proof_irrel), 34 | id_left := (λ a b f, sigma.eq !id_left !proof_irrel), 35 | id_right := (λ a b f, sigma.eq !id_right !proof_irrel) ⦄ 36 | 37 | -- ref:def:reduced-coslice 38 | -- Definition 4.2 39 | definition ReducedCoslice (C : Category) (c : C) := Mk (reduced_coslice C c) 40 | 41 | notation c `//` C := ReducedCoslice C c 42 | 43 | open functor 44 | 45 | definition forget (C : Category) (c : C) : (c // C) ⇒ C := 46 | ⦃ functor, 47 | object := λ a, to a, 48 | morphism := λ a b f, pr1 f, 49 | respect_id := λa, eq.refl _, 50 | respect_comp := λ a b c f g, eq.refl _ ⦄ 51 | 52 | end reduced_coslice 53 | 54 | open reduced_coslice 55 | open invcat 56 | 57 | open functor 58 | 59 | -- ref:def:matching-object 60 | -- Definition 4.4 61 | namespace matching_object 62 | 63 | open poly_unit reduced_coslice.red_coslice_obs reduced_coslice.coslice_obs 64 | 65 | definition matching_object.{u} {C : Category.{1 1}} [invcat C] (X : C ⇒ Type_category.{u}) (z : C) := 66 | -- we use an explicit definition of [limit_obj (limit_in_pretype (X ∘f (forget C z)))]: a natural transformation from the contstant functor 67 | Nat(𝟙, (X ∘f (forget C z))) 68 | 69 | definition matching_obj_map {C : Category.{1 1}} [invC : invcat C] (X : C ⇒ Type_category) (z : C) : 70 | X z → matching_object X z := 71 | begin 72 | intros x, unfold matching_object, unfold forget, unfold functor.compose, 73 | refine natural_transformation.mk (λ a u, X (hom_to a) x) _, 74 | { intros, esimp, rewrite id_right, cases f with [f_hom, tr], apply funext, intro y, 75 | esimp at *, 76 | unfold red_coslice_obs.to_coslice_obs at *, rewrite -tr, 77 | apply eq.symm, apply happly (respect_comp X f_hom (hom_to a)) x } 78 | end 79 | open natural_transformation 80 | 81 | definition nat_map_matching_obj_map {C : Category.{1 1}} [invC : invcat C] (X : C ⇒ Type_category) (z : C) (x : X z) : 82 | natural_map (matching_obj_map X z x) = (λ a u, X (hom_to a) x) := 83 | begin 84 | reflexivity 85 | end 86 | 87 | end matching_object 88 | 89 | namespace reedy 90 | universe variable u 91 | variables {C : Category.{1 1}} [invcat C] 92 | 93 | open matching_object 94 | 95 | -- ref:def:reedy-fibration 96 | -- Definition 4.6 (a Reedy fibrant diagram) 97 | definition is_reedy_fibrant [class] (X : C ⇒ U) := 98 | Π z, is_fibration (matching_obj_map X z) 99 | end reedy 100 | -------------------------------------------------------------------------------- /2ltt/matching_facts.lean: -------------------------------------------------------------------------------- 1 | import algebra.category 2 | import .inverse .limit .fibrant .finite .matching 3 | 4 | open sigma category eq.ops function functor 5 | 6 | open reduced_coslice invcat 7 | 8 | section MO_facts 9 | variables {C : Category.{1 1}} 10 | {D : Category} 11 | 12 | open reduced_coslice reduced_coslice.coslice_obs nat subcat_obj fincat matching_object invcat 13 | open equiv equiv.ops poly_unit 14 | open natural_transformation 15 | 16 | variables {z : C} {x : C_without_z z} 17 | {φ : C ⇒ ℕop} {reflecting_id : id_reflect φ} 18 | {max_rank : ∀ x, φ x ≤ φ z} 19 | [is_obj_finite C] [invcat C] 20 | (X : C ⇒ U) 21 | 22 | definition red_coslice_to_C' (o : x//C_without_z z) : (obj x)//C := 23 | begin 24 | cases o with [t,f,non_id], 25 | refine red_coslice_obs.mk (obj t) f _, intros p, intros q, 26 | assert H : @eq (C_without_z z) x t, begin cases x, cases t, esimp at *, congruence, assumption end, 27 | apply non_id H, cases H, esimp, apply q 28 | end 29 | 30 | definition red_coslice_to_C'_hom (a b : x//C_without_z z) (f : a ⟶ b) : (red_coslice_to_C' a) ⟶ (red_coslice_to_C' b) := 31 | begin 32 | cases f with [f, comm_tr], cases a with [a,g,non_id], cases b with [b,g', non_id'], 33 | unfold red_coslice_obs.to_coslice_obs at *, unfold red_coslice_to_C' at *, esimp at *, 34 | apply ⟨f,comm_tr⟩ 35 | end 36 | 37 | -- (Danil) I pack this property as a lemma to prevent unnessesary unfolding 38 | -- which sometimes leads to errors when unfolding definitions dependent on this one 39 | lemma subcat_obj_eq {c c1: C} {P : C → Prop} {p : P c} {p1 : P c1} 40 | (q : subcat_obj.mk c p = subcat_obj.mk c1 p1) : c = c1 := 41 | begin refine subcat_obj.no_confusion q (λ x y, x) end 42 | 43 | include φ max_rank reflecting_id 44 | 45 | definition red_coslice_ne_z (t : C) (f : x ⟶ t) : t ≠ z := 46 | begin 47 | cases x with [x', x'_ne], esimp at *, 48 | cases (@fincat.has_decidable_eq _ _ t z) with [t_eq_z, t_ne_z], cases t_eq_z, exfalso, 49 | apply @no_incoming_non_id_arrows _ z φ max_rank reflecting_id, existsi x', existsi f, apply x'_ne, apply t_ne_z 50 | end 51 | 52 | definition red_coslice_from_C' [reducible] (o : (obj x)//C) : x//C_without_z z := 53 | begin 54 | cases o with [t,f,non_id], esimp at *, 55 | cases x with [x', x'_ne], esimp at *, 56 | let t' := subcat_obj.mk t 57 | (@red_coslice_ne_z _ z (mk x' x'_ne) φ reflecting_id max_rank _ _ t f), 58 | refine red_coslice_obs.mk _ _ _, apply t', apply f, intros p, 59 | intros q, 60 | have q1 : x' = t, from subcat_obj_eq p, 61 | apply non_id q1, cases q1, esimp at *, apply q 62 | end 63 | 64 | definition red_coslice_from_C'_hom {a b : x//C} (f : a ⟶ b) : 65 | (@red_coslice_from_C' _ _ _ _ reflecting_id max_rank _ _ a) ⟶ (@red_coslice_from_C' _ _ _ _ reflecting_id max_rank _ _ b) := 66 | begin 67 | cases f with [f, comm_tr], cases a with [a,g,non_id], cases b with [b,g', non_id'], 68 | unfold red_coslice_obs.to_coslice_obs at *, unfold red_coslice_from_C' at *, esimp at *, 69 | cases x, esimp at *, 70 | apply ⟨f,comm_tr⟩, 71 | end 72 | 73 | definition red_coslice_without_z_equiv : (obj x)//C ≃ₛ (x//C_without_z z) := 74 | equiv.mk (@red_coslice_from_C' C _ _ _ reflecting_id max_rank _ _) red_coslice_to_C' 75 | begin 76 | intros y, cases y with [y,f,non_id], 77 | cases x with [x, x_ne], esimp 78 | end 79 | begin 80 | intros y, cases y with [y,f,non_id], unfold red_coslice_to_C', 81 | cases y, esimp, unfold red_coslice_from_C', cases x, esimp 82 | end 83 | 84 | definition red_coslice_eq_1 : 85 | ∀ y, object X (obj (red_coslice_obs.to ((@red_coslice_without_z_equiv _ _ x _ reflecting_id max_rank _ _) ∙ y))) = 86 | X (red_coslice_obs.to y) := 87 | begin 88 | intro, unfold fn, cases y with [y,f,y_ne], cases x with [x, x_ne], esimp 89 | end 90 | 91 | definition MO'_to_MO_map : 92 | matching_object (Functor_from_C' z X) x → matching_object X (obj x) := 93 | begin 94 | let ψ := @red_coslice_without_z_equiv C z _ φ reflecting_id max_rank _ _, 95 | intros a, 96 | refine natural_transformation.mk _ _, 97 | { intros y uu, unfold functor.compose at *, unfold forget at *, esimp at *, unfold red_coslice_obs.to_coslice_obs, 98 | unfold Functor_from_C' at *, 99 | have HH : object X (red_coslice_obs.to (ψ ∙ y)), from (natural_map a) (ψ ∙ y) star, 100 | intro, unfold fn at *, 101 | cases y with [y,f,y_ne], cases x with [x', x_ne], esimp at *, 102 | apply HH }, 103 | { cases a with [η, NatSq], intros x' y, 104 | intros f', 105 | let X' := Functor_from_C' z X, 106 | let C' := C_without_z z, 107 | assert HH : #function 108 | morphism (X' ∘f forget C' x) (red_coslice_from_C'_hom f') ∘ η (red_coslice_from_C' x') = 109 | η (@red_coslice_from_C' _ _ _ _ reflecting_id max_rank _ _ y), 110 | begin apply NatSq _ end, 111 | cases x with [x,f,non_id_f], cases y with [y,g,non_id_g], esimp at *, 112 | cases f' with [a,p1], esimp, 113 | cases x' with [o, o_ne], esimp, 114 | unfold Functor_from_C' at *, unfold functor.compose at *, unfold forget at *, esimp at *, 115 | unfold red_coslice_from_C'_hom at *, esimp at *, 116 | rewrite natural_map_proj, 117 | apply funext, intros uu, esimp, esimp, cases uu, refine happly HH _ } 118 | end 119 | 120 | definition MO_equiv : matching_object (Functor_from_C' z X) x ≃ₛ matching_object X (obj x) := 121 | begin 122 | let C' := C_without_z z, 123 | let X' := Functor_from_C' z X, 124 | let ψ := @red_coslice_without_z_equiv C z x φ reflecting_id max_rank _ _, 125 | 126 | -- we got two commuting triangles using equivalence ((obj x)//C) ≃ (x//C') 127 | assert Htr1 : ∀ y, (X' ∘f !forget) (@to_fun _ _ ψ y) = (X ∘f !forget) y, 128 | begin intro, unfold functor.compose, unfold forget, unfold Functor_from_C', 129 | unfold red_coslice_obs.to_coslice_obs, esimp, refine red_coslice_eq_1 _ _, 130 | end, 131 | assert Htr2 : ∀ y, (X ∘f !forget) (@inv_fun _ _ ψ y) = (X' ∘f (!forget)) y, 132 | begin intro, unfold functor.compose, unfold forget, unfold Functor_from_C', 133 | unfold red_coslice_obs.to_coslice_obs, cases y with [y,f,y_ne], esimp end, 134 | unfold matching_object, 135 | refine equiv.mk (@MO'_to_MO_map _ _ _ _ reflecting_id max_rank _ _ X) _ _ _, 136 | { intros a, cases a with [η, NatSq], refine natural_transformation.mk _ _, 137 | intros y uu, 138 | have η' : poly_unit → (X∘f forget C (obj x)) (@equiv.inv _ _ ψ y), from η _, 139 | intro, unfold functor.compose at *, unfold forget at *, unfold Functor_from_C' at *, 140 | unfold red_coslice_obs.to_coslice_obs, cases y with [y,f,y_ne], esimp, apply η' star, 141 | intros a b f, 142 | assert HH : morphism (X∘f forget C (obj x)) (red_coslice_to_C'_hom _ _ f) ∘ η (red_coslice_to_C' a) 143 | = η (red_coslice_to_C' b), 144 | begin refine NatSq _ end, 145 | unfold Functor_from_C' at *, unfold functor.compose at *, unfold forget at *, esimp at *, 146 | unfold red_coslice_to_C'_hom at *, esimp at *, 147 | cases f with [f,comm_tr], cases a with [a, ff, non_id_ff ], cases b with [b, gg, non_id_gg], 148 | apply funext, intros uu, cases uu, 149 | refine happly HH _}, 150 | { intros a, refine nat_trans_eq, cases a with [η, NatSq], rewrite natural_map_proj, 151 | unfold MO'_to_MO_map, unfold fn, unfold equiv.inv, rewrite natural_map_proj, 152 | apply funext, intro y, apply funext, intro uu, cases uu, 153 | have Heq : (@to_fun _ _ ψ (@inv_fun (obj x//C) _ ψ y)) = y, from by apply right_inv, 154 | cases Heq, cases y with [y,f,non_id_f], esimp at *, cases x with [x,p], esimp }, 155 | { intros a, refine nat_trans_eq, cases a with [η, NatSq], 156 | esimp, rewrite natural_map_proj, 157 | unfold MO'_to_MO_map, rewrite natural_map_proj, 158 | apply funext, intro y, apply funext, intro uu, cases uu, 159 | cases y, cases x, esimp } 160 | end 161 | 162 | end MO_facts 163 | -------------------------------------------------------------------------------- /2ltt/pullbacks.lean: -------------------------------------------------------------------------------- 1 | import .fibrant .limit data.equiv algebra.category .facts 2 | 3 | open function 4 | 5 | section pullback 6 | 7 | universe variable u 8 | variables {A B C : Type.{u}} 9 | (f : A → C) (g : B → C) 10 | {isfib : is_fibration f} 11 | 12 | structure Pullback (f : A → C) (g : B → C):= 13 | (pA : A) 14 | (pB : B) 15 | (pC : C) 16 | (p_law : (f pA = pC) × (g pB = pC)) 17 | 18 | definition Pullback' : Type := Σ (b : B), fibreₛ f (g b) 19 | 20 | open sigma.ops 21 | 22 | -- ref:lem:fib-closure:pb 23 | -- Lemma 3.9 (i) 24 | definition Pullback'_is_fibrant : 25 | is_fibration (λ (pb : Pullback' f g), pb.1) := 26 | λ b, @equiv_is_fibrant _ _ (equiv.symm (fibre_projection b)) (isfib (g b)) 27 | 28 | -- Inspired by the implementation from HoTT Lean library 29 | inductive pb_cat_ob : Type := 30 | | obA : pb_cat_ob 31 | | obB : pb_cat_ob 32 | | obC : pb_cat_ob 33 | 34 | 35 | open pb_cat_ob 36 | inductive pb_cat_hom : pb_cat_ob → pb_cat_ob → Type := 37 | | f1 : pb_cat_hom obA obC 38 | | f2 : pb_cat_hom obB obC 39 | 40 | open sum 41 | 42 | definition no_comp_pullback_hom {a b c : pb_cat_ob} 43 | : pb_cat_hom a b → pb_cat_hom b c → false := 44 | begin intros f g, cases g: cases f end 45 | 46 | 47 | -- Pullback as a category with three objects obA obB and obC 48 | -- and arrows f1 : obA → obC and f2: obB → obC 49 | definition pullback_category : category pb_cat_ob := 50 | ⦃ category, 51 | hom := λ a b, pb_cat_hom a b + (a = b), 52 | comp := λ a b c f g, 53 | begin 54 | cases f with [f₁, f₂]: cases g with [g₁, g₂], 55 | { apply inl, exfalso, apply no_comp_pullback_hom g₁ f₁}, 56 | { apply inl, rewrite g₂, assumption}, 57 | { apply inl, rewrite -f₂, assumption}, 58 | { apply inr, rewrite -f₂, assumption} 59 | end, 60 | ID := λ a, inr rfl , 61 | assoc := λ a b c d h g f, 62 | begin 63 | induction h with rh ph: induction g with rg pg: induction f with rf pf: 64 | try cases pg: try cases pf: esimp: exfalso: apply no_comp_pullback_hom; assumption; assumption 65 | end, 66 | id_left := λ a b f, begin cases f with [fl, fr], esimp, cases fr, esimp end, 67 | id_right := λ a b f, begin cases f with [fl, fr], esimp, cases fr, esimp end ⦄ 68 | 69 | open category 70 | 71 | definition PullbackCat := Mk pullback_category 72 | 73 | definition pullback_diagram (f : A → C) (g : B → C) : PullbackCat ⇒ Type_category.{u} := 74 | ⦃ functor, 75 | object := λ c, 76 | match c with 77 | | obA := A 78 | | obB := B 79 | | obC := C 80 | end, 81 | morphism := λ a b h, 82 | begin 83 | cases h with [pb_hom, p], 84 | cases pb_hom, 85 | apply f, apply g, cases p, apply id 86 | end, 87 | respect_id := λ a, begin reflexivity end, 88 | respect_comp := 89 | begin 90 | intros a b c pf pg, 91 | cases pf with [pb_hom_f, id_f]: 92 | cases pg with [pb_hom_g, id_g]: 93 | try (cases pb_hom_f: cases pb_hom_g: esimp): 94 | try cases id_g: esimp: try cases id_f: esimp 95 | end ⦄ 96 | 97 | open poly_unit 98 | 99 | open equiv 100 | 101 | definition const_funct_unit [reducible] [unfold_full] := 102 | const_funct_obj PullbackCat Type_category poly_unit 103 | 104 | definition nat_unit_Pullback_equiv : 105 | const_funct_unit ⟹ pullback_diagram f g ≃ₛ Pullback f g:= 106 | begin 107 | esimp, 108 | refine (equiv.mk _ _ _ _), 109 | { intro N, cases N with [η, NatSq], unfold pullback_diagram at *, esimp at *, 110 | refine Pullback.mk _ _ _ _, 111 | apply (η obA star), apply (η obB star), apply (η obC star), esimp, 112 | refine (_,_), 113 | have H : f ∘ η obA = η obC ∘ id, from NatSq (inl pb_cat_hom.f1), apply happly H, 114 | have H : g ∘ η obB = η obC ∘ id, from NatSq (inl pb_cat_hom.f2), apply happly H }, 115 | { intro PB, cases PB, cases p_law with [f_eq, g_eq], 116 | refine natural_transformation.mk _ _, 117 | intros a, unfold pullback_diagram, esimp, intro uu, cases a: esimp: assumption, 118 | intros a b pb, unfold pullback_diagram, cases pb with [pb_hom, p_id], 119 | cases pb_hom: esimp, 120 | { apply funext, intros, apply f_eq}, 121 | { apply funext, intros, apply g_eq}, 122 | { cases p_id, esimp}}, 123 | { unfold function.left_inverse, intros x, cases x, 124 | esimp at *, congruence, apply funext, intros, apply funext, intro uu, esimp, 125 | cases x: esimp: cases uu: reflexivity }, 126 | { unfold function.right_inverse, unfold function.left_inverse, intros x, 127 | cases x, cases p_law with [f_eq, g_eq], esimp } 128 | end 129 | 130 | definition Pullback'_Pullback_equiv : Pullback' f g ≃ₛ Pullback f g:= 131 | equiv.mk 132 | (λ x, begin cases x with [p₁, p₂], cases p₂ with [pp₁, pp₂], 133 | apply Pullback.mk pp₁ p₁ (g p₁) (pp₂, rfl) end) 134 | (λ x, begin cases x, cases p_law with [p₁, p₂], refine ⟨ pB, #eq.ops ⟨pA, p₁ ⬝ p₂⁻¹⟩⟩end) 135 | (λ x, begin cases x with [p₁, p₂], cases p₂ with [pp₁, pp₂], esimp end) 136 | (λ x, begin cases x, cases p_law with [p₁, p₂], esimp, cases p₂, esimp end) 137 | 138 | end pullback 139 | -------------------------------------------------------------------------------- /2ltt/simplicial.lean: -------------------------------------------------------------------------------- 1 | import data.fin algebra.monotone algebra.category 2 | open function nat fin 3 | 4 | definition fin_strict_order [instance] (n : ℕ) : strict_order (fin n) 5 | := strict_order.mk (λ i i', val i < val i') 6 | (λ i, lt.irrefl (val i)) 7 | (λ a b c, lt.trans) 8 | 9 | structure delta_plus (n m : ℕ) := 10 | (f : fin (succ n) -> fin (succ m)) 11 | (mon : strictly_increasing f) 12 | 13 | attribute delta_plus.f [coercion] 14 | 15 | definition id_delta_plus (n : ℕ) : delta_plus n n := 16 | delta_plus.mk id strictly_increasing_id 17 | 18 | definition comp_delta_plus (n m p : ℕ) 19 | (f : delta_plus m p)(g : delta_plus n m) : delta_plus n p := 20 | delta_plus.mk (f ∘ g) 21 | (strictly_increasing_comp_inc_inc (delta_plus.mon f) (delta_plus.mon g)) 22 | 23 | definition eq_delta_plus {n m : ℕ} 24 | {f g : delta_plus n m} (q : Π (i : fin (succ n)), f i = g i) : @eq (delta_plus n m) f g := 25 | begin cases f, cases g, esimp at *, congruence, apply funext, apply q end 26 | 27 | definition delta_plus_cat : category ℕ := 28 | ⦃ category, 29 | hom := delta_plus, 30 | comp := comp_delta_plus, 31 | ID := id_delta_plus, 32 | assoc := λ a b c d h g f, eq_delta_plus (λ i, rfl), 33 | id_right := λ a b f, eq_delta_plus (λ i, rfl), 34 | id_left := λ a b f, eq_delta_plus (λ i, rfl) ⦄ 35 | 36 | definition delta_plus_Cat : Category := category.Mk delta_plus_cat 37 | 38 | notation `Δ+` := delta_plus_Cat 39 | -------------------------------------------------------------------------------- /2ltt/types/prod.lean: -------------------------------------------------------------------------------- 1 | import ..fibrant 2 | 3 | open prod prod.ops 4 | open fib_eq 5 | 6 | /- Ported from the Lean HoTT library, trying to keep as close as possible 7 | to the original implementation. 8 | Changes while porting: change equality notation (from "=" to "~"), 9 | replace usages of "reflexivity" with "simp" tactic. 10 | The "simp" tactic allows to apply computation rules for fibrant equality, which doesn't 11 | hold judgementally. 12 | The "induction" tactic applies fib_eq.elim induction principle automatically, as it is 13 | marked with [recursor] attribute. 14 | -/ 15 | 16 | namespace hprod 17 | 18 | open fib_eq.ap 19 | attribute trans [reducible] 20 | attribute ap [reducible] 21 | 22 | variables {A : Fib} 23 | variables {B : Fib} 24 | variables {u v w : A × B} {a a' : A} {b b' : B} 25 | {P Q : A → Fib} 26 | 27 | protected definition eta [unfold 3] (u : A × B) : (pr₁ u, pr₂ u) =ᵒ u := 28 | by cases u; apply idp 29 | 30 | definition pair_eq {a a' : A} {b b' : B} : (a =ᵒ a') × (b =ᵒ b') -> (a,b) =ᵒ (a',b') 31 | | pair_eq (p, q) := ap₂ (λ x y, (x,y)) p q 32 | 33 | 34 | definition pair_eq' [reducible] (pa : a =ᵒ a') (pb : b =ᵒ b') : (a, b) =ᵒ (a', b') := 35 | ap₂ (λ x y, (x,y)) pa pb 36 | 37 | definition pair_eq'_β {a a' : A} : pair_eq' (reflᵒ a) (reflᵒ a') = reflᵒ (a,a') := 38 | by unfold pair_eq'; apply ap₂_β 39 | 40 | definition prod_eq' (H₁ : u.1 =ᵒ v.1) (H₂ : u.2 =ᵒ v.2) : u =ᵒ v := 41 | by cases u; cases v; exact pair_eq' H₁ H₂ 42 | 43 | definition prod_eq [reducible] (p : pr₁ u =ᵒ pr₁ v) (q : pr₂ u =ᵒ pr₂ v) : u =ᵒ v := 44 | (prod.cases_on u (λ x y, prod.cases_on v (λ x1 y1, pair_eq'))) p q 45 | 46 | definition prod_eq_β : prod_eq (reflᵒ (pr₁ u)) (reflᵒ (pr₂ u)) = reflᵒ u := 47 | by cases u; unfold prod_eq; unfold pair_eq'; apply ap₂_β 48 | 49 | definition prod_eq_β' [simp] : prod_eq (reflᵒ a) (reflᵒ b) = reflᵒ (a,b) := 50 | by unfold prod_eq; unfold pair_eq'; rewrite ap₂_β 51 | 52 | definition eq_pr1 [reducible] (p : u =ᵒ v) : u.1 =ᵒ v.1 := 53 | ap pr1 p 54 | 55 | definition eq_pr2 [reducible] (p : u =ᵒ v) : u.2 =ᵒ v.2 := 56 | ap pr2 p 57 | 58 | namespace ops 59 | postfix `..1`:(max+1) := eq_pr1 60 | postfix `..2`:(max+1) := eq_pr2 61 | end ops 62 | 63 | open ops 64 | 65 | protected definition ap_pr1 (p : u =ᵒ v) : ap pr1 p =ᵒ p..1 := idp 66 | protected definition ap_pr2 (p : u =ᵒ v) : ap pr2 p =ᵒ p..2 := idp 67 | 68 | open fib_eq.ap 69 | 70 | definition pair_prod_eq (p : u.1 =ᵒ v.1) (q : u.2 =ᵒ v.2) : 71 | ((prod_eq p q)..1, (prod_eq p q)..2) =ᵒ (p, q) := 72 | by 73 | induction u; induction v; esimp at *; 74 | induction p; induction q; 75 | rewrite prod_eq_β'; repeat rewrite elimo_β 76 | 77 | /- Transport -/ 78 | 79 | definition prod_transport (p : a =ᵒ a') (u : P a × Q a) : 80 | p ▹ u =ᵒ (p ▹ u.1, p ▹ u.2) := 81 | by induction p; induction u; simp 82 | 83 | end hprod 84 | -------------------------------------------------------------------------------- /2ltt/types/sigma.lean: -------------------------------------------------------------------------------- 1 | import ..fibrant ..facts 2 | /- Ported from the Lean HoTT library with some changes. 3 | Facts about sigma-types harder to port in a way they are implemented. 4 | Proofs use more "generic" functionality from the HoTT library in comparison 5 | with product types. 6 | -/ 7 | open sigma sigma.ops 8 | 9 | namespace hsigma 10 | 11 | open fib_eq fib_eq.ap 12 | attribute elimo_β [simp] 13 | attribute trans [reducible] 14 | 15 | variables {A : Fib} 16 | variables {B : Fib} 17 | {P Q : A → Fib} 18 | {P' Q' : B → Fib} 19 | variables {u v w : Σ (a: A), P a} {a a' : A} {b b₁ b₂ : P a} {b' : P a'} 20 | {u' v' w' : Σ (b : B), P' b} 21 | 22 | definition dpair_eq_dpair [reducible] (p : a =ᵒ a') (q : p ▹ b =ᵒ b') : ⟨a, b⟩ =ᵒ ⟨a', b'⟩ := 23 | begin 24 | induction p, rewrite transport_β at q, 25 | induction q, reflexivity 26 | end 27 | 28 | -- it is awkward to define such computation rules, because we have to 29 | -- manually use other computation rules in the statement itself 30 | -- But it seems, that it is sufficient to define computation rules for 31 | -- some base cases, and then just unfold definitions and use "simp" or 32 | -- rewrite explicitly. 33 | 34 | definition dpair_eq_dpair_β : 35 | @dpair_eq_dpair _ _ a a _ b idp (eq.symm (transport_β b) ▹s idp) = idp := 36 | begin unfold dpair_eq_dpair, rewrite elimo_β, 37 | rewrite eq.transport_concat, simp end 38 | 39 | definition sigma_eq [reducible] (p : pr₁ u =ᵒ pr₁ v) (q : p ▹ pr₂ u =ᵒ pr₂ v) : u =ᵒ v := 40 | (sigma.cases_on u (λ x y, sigma.cases_on v (λ x1 y1, dpair_eq_dpair))) p q 41 | 42 | definition sigma_to_dpair_eq (u v : Σ a, P a) : 43 | u =ᵒ v -> Σ (p : pr₁ u =ᵒ pr₁ v), p ▹ (pr₂ u) =ᵒ pr₂ v := 44 | elimᵒ ⟨reflᵒ u.1, eq.rec (reflᵒ u.2) (eq.symm (transport_β u.2))⟩ _ 45 | 46 | namespace fib_eq 47 | definition pathover_idp_of_eq [reducible] {A : Fib} {B : A → Fib} {a : A} {b b' : B a} : 48 | b =ᵒ b' → (reflᵒ a) ▹ b =ᵒ b' := λ p, by induction p; rewrite transport_β 49 | 50 | definition pathover_idp_of_eq_β {A : Fib} {B : A → Fib} {a : A} {b : B a} : 51 | @pathover_idp_of_eq A _ _ _ _ (reflᵒ b) = (eq.symm (transport_β b) ▹s idp) := 52 | by simp 53 | 54 | end fib_eq 55 | 56 | /- Applying dpair to one argument is the same as dpair_eq_dpair with reflexivity in the first place. -/ 57 | 58 | open fib_eq.ap 59 | 60 | definition ap_dpair (q : b₁ =ᵒ b₂) : 61 | ap (sigma.mk a) q =ᵒ dpair_eq_dpair idp (fib_eq.pathover_idp_of_eq q) := 62 | begin 63 | induction q, unfold fib_eq.pathover_idp_of_eq, unfold dpair_eq_dpair, 64 | repeat rewrite elimo_β, rewrite eq.transport_concat 65 | end 66 | 67 | -- proof of the same property as above, but using propositional computation rules for 68 | -- the definitions involved in the type 69 | definition ap_dpair' (q : b₁ =ᵒ b₂) : 70 | ap (sigma.mk a) q =ᵒ dpair_eq_dpair idp (fib_eq.pathover_idp_of_eq q) := 71 | begin 72 | induction q, rewrite fib_eq.pathover_idp_of_eq_β, rewrite dpair_eq_dpair_β, 73 | rewrite ap_β 74 | end 75 | 76 | end hsigma 77 | -------------------------------------------------------------------------------- /Dockerfile: -------------------------------------------------------------------------------- 1 | FROM ubuntu:18.04 2 | MAINTAINER Danil Annenkov 3 | 4 | # Lean build flags 5 | ENV CMAKE_FLAGS="-D CMAKE_BUILD_TYPE=RELEASE -D BOOST=OFF -D TCMALLOC=OFF -G Ninja" 6 | 7 | # Set non-interactive mode 8 | ARG DEBIAN_FRONTEND=noninteractive 9 | 10 | ENV LANG C.UTF-8 11 | 12 | # Install the required packages 13 | 14 | RUN apt-get update && \ 15 | apt-get install \ 16 | nano \ 17 | python \ 18 | git \ 19 | libgmp-dev \ 20 | libmpfr-dev \ 21 | emacs \ 22 | cmake \ 23 | liblua5.2.0 \ 24 | lua5.2-0 \ 25 | lua5.2-dev \ 26 | gitg \ 27 | ninja-build \ 28 | valgrind \ 29 | doxygen -y 30 | 31 | # RUN update-alternatives --remove-all gcc 32 | 33 | # RUN update-alternatives --remove-all g++ 34 | 35 | # RUN apt-get update 36 | RUN apt-get install g++-4.8 -y 37 | RUN apt-get upgrade -y && apt-get dist-upgrade -y 38 | 39 | RUN update-alternatives --install /usr/bin/g++ g++ /usr/bin/g++-4.8 0 40 | 41 | RUN g++ --version 42 | 43 | # Download from the repository, build, clean up 44 | RUN \ 45 | git clone --depth 1 https://github.com/leanprover/lean2.git && \ 46 | mkdir -p lean2/build && \ 47 | (cd lean2/build; cmake $CMAKE_FLAGS ../src && ninja && ninja install) && \ 48 | lean --version && \ 49 | rm -rf lean 50 | 51 | RUN mkdir /root/.emacs.d 52 | ADD extra/init.el /root/.emacs.d 53 | RUN ls ~/.emacs.d 54 | RUN emacs --batch -l /root/.emacs.d/init.el 55 | 56 | RUN mkdir /root/2ltt 57 | ADD 2ltt /root/2ltt/ 58 | 59 | RUN cd ~/2ltt && linja 60 | 61 | # This is just a convenience for running the container. It can be overridden. 62 | ENTRYPOINT ["/bin/bash"] 63 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | all: 2 | $(MAKE) -C 2ltt 3 | 4 | docker: 5 | docker build -t lean-image . 6 | 7 | run-image: 8 | docker run -it lean-image 9 | 10 | clean: 11 | $(MAKE) clean -C 2ltt 12 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Lean Formalisation of Two-Level Type Theory 2 | 3 | Files in the `2ltt` directory: 4 | 5 | | File | | 6 | |--------------------|:--------------------------------------------------------------------: 7 | | fibrant.lean | implementation of the two-level type theory 8 | | finite.lean | facts about finite sets and categories with a finite set of objects 9 | | limit.lean | definition of limits and construction of limits in category of pretypes 10 | | inverse.lean | definition of inverse categories (Section 4.2) 11 | | pullbacks.lean | definition of a pullback, constructed explicitly and using the limit of a diagram along with a proof these definitions are equivalent. Proof of that fibrations are closed under pullbacks. 12 | | matching.lean | definition of the matching object (Section 4.3) 13 | | matching_facts.lean | facts about the matching object from the category C with one object removed 14 | | fibrantlimits_aux.lean | auxiliary lemmas for the proof of the fibrant limits theorem including equivalences forming the core of the proof 15 | | fibrantlimits.lean | a proof of the fibrant limit theorem (Section 4.3, Theorem 4.8.) 16 | | simplicial.lean | initial definition of semi-simplicial types (work in progress) 17 | | facts.lean | some auxiliary lemmas which we could not find in the standard library 18 | | types/* | some examples of reasoning in the inner (fibrant) theory. Mainly, ported from the Lean 2 HoTT library. 19 | 20 | ## Compilation 21 | 22 | Requires Lean 2 (version 0.2.0) to compile. 23 | 24 | Lean 2 installation instructions in Ubuntu : https://github.com/leanprover/lean2/blob/master/doc/make/ubuntu-12.04-detailed.md 25 | 26 | After installing Lean 2, in the terminal, go to the directory containing the development and use ```make``` to compile the project, or run go to the `2ltt` folder and run ```linja``` directly. 27 | 28 | Use your favorite editor to navigate the source code (see some hints about emacs lean-mode blow). 29 | 30 | ### Using Docker 31 | 32 | Alternatively, use `Docker` image. The image is based on Ubuntu and contains Lean 2, our development and Emacs with `lean-mode`, allowing to navigate through the development and compile it. 33 | 34 | Steps to build a Docker image: 35 | * In the terminal, go to the directory containing the development and run `make docker` 36 | * After building step is complete (this might take some time), run `make run-image` to connect your current terminal session to the container. 37 | * In the container, type `emacs /root/2ltt/fibrantlimits.lean` to start editing `fibrantlimits.lean` with the lean-mode in Emacs. 38 | 39 | 40 | ### Interacting with the Lean development in the Docker image 41 | 42 | As we have said before, the docker image comes with Emacs and lean-mode. The lean-mode provides syntax highlighting and allows to interactively develop proofs. 43 | Use C-c C-g to show goal in tactic proof, C-c C-p to print information about identifier, C-c C-t to show type, and C-c C-l to comple a lean file. 44 | 45 | Alternatively, one can use the command line to compile the project or individual files: 46 | 47 | * After executing `make run-image`, change directory (in the Ubuntu's command line) to our Lean development `cd /root/2ltt`. 48 | * Run `make clean` to clean previously compiled files 49 | * Run `make` again to rebuild the project. 50 | * Individual files can be compiled using `lean `. 51 | 52 | The image comes with the preinstalled `nano` editor that can be used to view/edit the source files (but without syntax highlighting). Other editors can be installed, if required, using the `apt-get` command. For example, `apt-get install vim`. 53 | 54 | -------------------------------------------------------------------------------- /extra/init.el: -------------------------------------------------------------------------------- 1 | (require 'package) 2 | (add-to-list 'package-archives '("melpa" . "http://melpa.milkbox.net/packages/") t) 3 | (when (< emacs-major-version 24) 4 | ;; For important compatibility libraries like cl-lib 5 | (add-to-list 'package-archives '("gnu" . "http://elpa.gnu.org/packages/"))) 6 | (package-initialize) 7 | 8 | ;; Install required/optional packages for lean-mode 9 | (defvar lean-mode-required-packages 10 | '(company dash dash-functional flycheck f 11 | fill-column-indicator s)) 12 | (let ((need-to-refresh t)) 13 | (dolist (p lean-mode-required-packages) 14 | (when (not (package-installed-p p)) 15 | (when need-to-refresh 16 | (package-refresh-contents) 17 | (setq need-to-refresh nil)) 18 | (package-install p)))) 19 | 20 | ;; Set up lean-root path 21 | (setq lean-rootdir "/lean2") 22 | (setq-local lean-emacs-path 23 | (concat (file-name-as-directory lean-rootdir) 24 | (file-name-as-directory "src") 25 | "emacs")) 26 | (add-to-list 'load-path (expand-file-name lean-emacs-path)) 27 | (require 'lean-mode) -------------------------------------------------------------------------------- /release.sh: -------------------------------------------------------------------------------- 1 | #!/bin/bash 2 | git archive -o code.zip HEAD 3 | --------------------------------------------------------------------------------