prove that iota is n-truncated/n-connected if the maps in the sequence are
This commit is contained in:
parent
1a35543661
commit
899e3cf2e4
4 changed files with 112 additions and 28 deletions
|
@ -110,4 +110,4 @@ These projects are done
|
|||
- Installation instructions for Lean 2 can be found [here](https://github.com/leanprover/lean2).
|
||||
- Some notes on the Emacs mode can be found [here](https://github.com/leanprover/lean2/blob/master/src/emacs/README.md) (for example if some unicode characters don't show up, or increase the spacing between lines by a lot).
|
||||
- If you contribute, please use rebase instead of merge (e.g. `git pull -r`).
|
||||
- We try to separate the repository into the folders `algebra`, `homotopy`, `homology` and `cohomology`. Homotopy theotic properties of types which do not explicitly mention homotopy, homology or cohomology groups (such as `A ∧ B ≃* B ∧ A`) are part of `homotopy`.
|
||||
- We try to separate the repository into the folders `algebra`, `homotopy`, `homology`, `cohomology` and `colimit`. Homotopy theotic properties of types which do not explicitly mention homotopy, homology or cohomology groups (such as `A ∧ B ≃* B ∧ A`) are part of `homotopy`.
|
|
@ -14,7 +14,12 @@ namespace seq_colim
|
|||
abbreviation ι [constructor] := @inclusion
|
||||
abbreviation ι' [constructor] [parsing_only] {A} (f n) := @inclusion A f n
|
||||
|
||||
universe variable v
|
||||
variables {A A' A'' : ℕ → Type} (f : seq_diagram A) (f' : seq_diagram A') (f'' : seq_diagram A'')
|
||||
(τ τ₂ : Π⦃n⦄, A n → A' n) (p : Π⦃n⦄ (a : A n), τ (f a) = f' (τ a))
|
||||
(p₂ : Π⦃n⦄ (a : A n), τ₂ (f a) = f' (τ₂ a))
|
||||
(τ' : Π⦃n⦄, A' n → A'' n) (p' : Π⦃n⦄ (a' : A' n), τ' (f' a') = f'' (τ' a'))
|
||||
{P : Π⦃n⦄, A n → Type.{v}} (g : seq_diagram_over f P) {n : ℕ} {a : A n}
|
||||
|
||||
definition lrep_glue {n m : ℕ} (H : n ≤ m) (a : A n) : ι f (lrep f H a) = ι f a :=
|
||||
begin
|
||||
|
@ -50,7 +55,7 @@ begin
|
|||
apply eq_pathover_id_right,
|
||||
refine (ap_compose (ι f) (colim_back f) _) ⬝ph _,
|
||||
refine ap02 _ _ ⬝ph _, rotate 1,
|
||||
{ rexact elim_glue f _ _ a},
|
||||
{ rexact elim_glue f _ _ a },
|
||||
refine _ ⬝pv ((natural_square (lrep_glue f (zero_le k))
|
||||
(ap (lrep_back f (zero_le k)) (left_inv (@f k) a)))⁻¹ʰ ⬝h _),
|
||||
{ exact (glue f _)⁻¹ ⬝ ap (ι f) (right_inv (lrep f (zero_le (succ k))) (f a)) },
|
||||
|
@ -62,8 +67,7 @@ begin
|
|||
refine ap02 _ (whisker_left _ (adj (@f _) a)) ⬝pv _,
|
||||
rewrite [-+ap_con, -ap_compose', ap_id],
|
||||
apply natural_square_tr },
|
||||
intro a,
|
||||
reflexivity,
|
||||
{ intro a, reflexivity }
|
||||
end
|
||||
|
||||
definition equiv_of_is_equiseq [constructor] (H : is_equiseq f) : seq_colim f ≃ A 0 :=
|
||||
|
@ -73,7 +77,6 @@ variable (f)
|
|||
end
|
||||
|
||||
section
|
||||
variables {n : ℕ} (a : A n)
|
||||
|
||||
definition rep_glue (k : ℕ) (a : A n) : ι f (rep f k a) = ι f a :=
|
||||
begin
|
||||
|
@ -85,9 +88,6 @@ end
|
|||
/- functorial action and equivalences -/
|
||||
section functor
|
||||
variables {f f' f''}
|
||||
variables (τ τ₂ : Π⦃n⦄, A n → A' n) (p : Π⦃n⦄ (a : A n), τ (f a) = f' (τ a))
|
||||
(p₂ : Π⦃n⦄ (a : A n), τ₂ (f a) = f' (τ₂ a))
|
||||
(τ' : Π⦃n⦄, A' n → A'' n) (p' : Π⦃n⦄ (a' : A' n), τ' (f' a') = f'' (τ' a'))
|
||||
include p
|
||||
definition seq_colim_functor [unfold 7] : seq_colim f → seq_colim f' :=
|
||||
begin
|
||||
|
@ -334,8 +334,6 @@ kshift_equiv f n ⬝e equiv_of_is_equiseq (λk, H (n+k) !le_add_right)
|
|||
|
||||
section over
|
||||
|
||||
universe variable v
|
||||
variables {P : Π⦃n⦄, A n → Type.{v}} (g : seq_diagram_over f P) {n : ℕ} {a : A n}
|
||||
variable {f}
|
||||
|
||||
definition rep_f_equiv_natural {k : ℕ} (p : P (rep f k (f a))) :
|
||||
|
@ -809,6 +807,10 @@ equiv.MK (seq_colim_trunc_of_trunc_seq_colim f k) (trunc_seq_colim_of_seq_colim_
|
|||
refine !ap_compose'⁻¹ ⬝ !elim_glue }
|
||||
end end
|
||||
|
||||
theorem is_conn_seq_colim [instance] (k : ℕ₋₂) [H : Πn, is_conn k (A n)] :
|
||||
is_conn k (seq_colim f) :=
|
||||
is_trunc_equiv_closed_rev -2 (trunc_seq_colim_equiv f k)
|
||||
|
||||
/- the colimit of a sequence of fibers is the fiber of the functorial action of the colimit -/
|
||||
definition domain_seq_colim_functor {A A' : ℕ → Type} {f : seq_diagram A}
|
||||
{f' : seq_diagram A'} (τ : Πn, A' n → A n)
|
||||
|
@ -828,7 +830,8 @@ begin
|
|||
refine _ ⬝e fiber_pr1 (seq_colim_over (seq_diagram_over_fiber τ p)) (ι f a),
|
||||
apply fiber_equiv_of_triangle (domain_seq_colim_functor τ p)⁻¹ᵉ,
|
||||
refine _ ⬝hty λx, (colim_sigma_triangle _ _)⁻¹,
|
||||
apply homotopy_inv_of_homotopy_pre (seq_colim_equiv _ _) (seq_colim_functor _ _) (seq_colim_functor _ _),
|
||||
apply homotopy_inv_of_homotopy_pre (seq_colim_equiv _ _)
|
||||
(seq_colim_functor _ _) (seq_colim_functor _ _),
|
||||
refine (λx, !seq_colim_functor_compose⁻¹) ⬝hty _,
|
||||
refine seq_colim_functor_homotopy _ _,
|
||||
intro n x, exact point_eq x.2,
|
||||
|
@ -840,8 +843,50 @@ definition fiber_seq_colim_functor0 {A A' : ℕ → Type} {f : seq_diagram A}
|
|||
{f' : seq_diagram A'} (τ : Πn, A' n → A n)
|
||||
(p : Π⦃n⦄, τ (n+1) ∘ @f' n ~ @f n ∘ @τ n) (a : A 0) :
|
||||
fiber (seq_colim_functor τ p) (ι f a) ≃ seq_colim (seq_diagram_fiber0 τ p a) :=
|
||||
sorry
|
||||
fiber_seq_colim_functor τ p a ⬝e
|
||||
seq_colim_equiv
|
||||
(λn, equiv_apd011 (λx y, fiber (τ x) y) (rep_pathover_rep0 f a))
|
||||
(λn x, sorry)
|
||||
-- maybe use fn_tro_eq_tro_fn2
|
||||
|
||||
variables {f f'}
|
||||
definition fiber_inclusion (x : seq_colim f) :
|
||||
fiber (ι' f 0) x ≃ fiber (seq_colim_functor (rep0 f) (λn a, idp)) x :=
|
||||
fiber_equiv_of_triangle (seq_colim_constant_seq (A 0))⁻¹ᵉ homotopy.rfl
|
||||
|
||||
theorem is_trunc_fun_seq_colim_functor (k : ℕ₋₂) (H : Πn, is_trunc_fun k (@τ n)) :
|
||||
is_trunc_fun k (seq_colim_functor τ p) :=
|
||||
begin
|
||||
intro x, induction x using seq_colim.rec_prop,
|
||||
exact is_trunc_equiv_closed_rev k (fiber_seq_colim_functor τ p a),
|
||||
end
|
||||
|
||||
open is_conn
|
||||
theorem is_conn_fun_seq_colim_functor (k : ℕ₋₂) (H : Πn, is_conn_fun k (@τ n)) :
|
||||
is_conn_fun k (seq_colim_functor τ p) :=
|
||||
begin
|
||||
intro x, induction x using seq_colim.rec_prop,
|
||||
exact is_conn_equiv_closed_rev k (fiber_seq_colim_functor τ p a) _
|
||||
end
|
||||
|
||||
variables (f f')
|
||||
theorem is_trunc_fun_inclusion (k : ℕ₋₂) (H : Πn, is_trunc_fun k (@f n)) :
|
||||
is_trunc_fun k (ι' f 0) :=
|
||||
begin
|
||||
intro x,
|
||||
apply @(is_trunc_equiv_closed_rev k (fiber_inclusion x)),
|
||||
apply is_trunc_fun_seq_colim_functor,
|
||||
intro n, apply is_trunc_fun_lrep, exact H
|
||||
end
|
||||
|
||||
theorem is_conn_fun_inclusion (k : ℕ₋₂) (H : Πn, is_conn_fun k (@f n)) :
|
||||
is_conn_fun k (ι' f 0) :=
|
||||
begin
|
||||
intro x,
|
||||
apply is_conn_equiv_closed_rev k (fiber_inclusion x),
|
||||
apply is_conn_fun_seq_colim_functor,
|
||||
intro n, apply is_conn_fun_lrep, exact H
|
||||
end
|
||||
|
||||
/- the sequential colimit of standard finite types is ℕ -/
|
||||
open fin
|
||||
|
|
|
@ -6,7 +6,7 @@ Authors: Floris van Doorn, Egbert Rijke
|
|||
|
||||
import ..move_to_lib types.fin types.trunc
|
||||
|
||||
open nat eq equiv sigma sigma.ops is_equiv is_trunc trunc prod fiber
|
||||
open nat eq equiv sigma sigma.ops is_equiv is_trunc trunc prod fiber function
|
||||
|
||||
namespace seq_colim
|
||||
|
||||
|
@ -28,14 +28,14 @@ namespace seq_colim
|
|||
attribute Seq_diagram.carrier [coercion]
|
||||
attribute Seq_diagram.struct [coercion]
|
||||
|
||||
variables {A A' : ℕ → Type} (f : seq_diagram A) (f' : seq_diagram A')
|
||||
variables {A A' : ℕ → Type} (f : seq_diagram A) (f' : seq_diagram A') {n m k : ℕ}
|
||||
include f
|
||||
|
||||
definition lrep {n m : ℕ} (H : n ≤ m) (a : A n) : A m :=
|
||||
definition lrep {n m : ℕ} (H : n ≤ m) : A n → A m :=
|
||||
begin
|
||||
induction H with m H y,
|
||||
{ exact a },
|
||||
{ exact f y }
|
||||
induction H with m H fs,
|
||||
{ exact id },
|
||||
{ exact @f m ∘ fs }
|
||||
end
|
||||
|
||||
definition lrep_irrel_pathover {n m m' : ℕ} (H₁ : n ≤ m) (H₂ : n ≤ m') (p : m = m') (a : A n) :
|
||||
|
@ -85,7 +85,6 @@ namespace seq_colim
|
|||
(lrep f H)⁻¹ᶠ
|
||||
|
||||
section generalized_lrep
|
||||
variable {n : ℕ}
|
||||
|
||||
-- definition lrep_pathover_lrep0 [reducible] (k : ℕ) (a : A 0) : lrep f k a =[nat.zero_add k] lrep0 f k a :=
|
||||
-- begin induction k with k p, constructor, exact pathover_ap A succ (apo f p) end
|
||||
|
@ -128,6 +127,9 @@ namespace seq_colim
|
|||
definition rep0 (m : ℕ) (a : A 0) : A m :=
|
||||
lrep f (zero_le m) a
|
||||
|
||||
definition rep_pathover_rep0 {n : ℕ} (a : A 0) : rep f n a =[nat.zero_add n] rep0 f n a :=
|
||||
!lrep_irrel_pathover
|
||||
|
||||
-- definition old_rep (m : ℕ) (a : A n) : A (n + m) :=
|
||||
-- by induction m with m y; exact a; exact f y
|
||||
|
||||
|
@ -150,6 +152,17 @@ namespace seq_colim
|
|||
{ apply pathover_ap, exact apo f IH}
|
||||
end
|
||||
|
||||
variable {f}
|
||||
definition is_trunc_fun_lrep (k : ℕ₋₂) (H : n ≤ m) (H2 : Πn, is_trunc_fun k (@f n)) :
|
||||
is_trunc_fun k (lrep f H) :=
|
||||
begin induction H with m H IH, apply is_trunc_fun_id, exact is_trunc_fun_compose k (H2 m) IH end
|
||||
|
||||
definition is_conn_fun_lrep (k : ℕ₋₂) (H : n ≤ m) (H2 : Πn, is_conn_fun k (@f n)) :
|
||||
is_conn_fun k (lrep f H) :=
|
||||
begin induction H with m H IH, apply is_conn_fun_id, exact is_conn_fun_compose k (H2 m) IH end
|
||||
|
||||
variable (f)
|
||||
|
||||
end generalized_lrep
|
||||
|
||||
section shift
|
||||
|
@ -228,7 +241,7 @@ namespace seq_colim
|
|||
definition seq_diagram_sigma [unfold 6] : seq_diagram (λn, Σ(x : A n), P x) :=
|
||||
λn v, ⟨f v.1, g v.2⟩
|
||||
|
||||
variables {n : ℕ} (f P)
|
||||
variables (f P)
|
||||
|
||||
theorem rep_f_equiv [constructor] (a : A n) (k : ℕ) :
|
||||
P (lrep f (le_add_right (succ n) k) (f a)) ≃ P (lrep f (le_add_right n (succ k)) a) :=
|
||||
|
@ -247,17 +260,17 @@ namespace seq_colim
|
|||
λn f x, g (f x)
|
||||
|
||||
variables {f f'}
|
||||
definition seq_diagram_over_fiber (g : Π⦃n⦄, A n → A' n)
|
||||
(p : Π⦃n⦄ (a : A n), g (f a) = f' (g a)) : seq_diagram_over f' (λn, fiber (@g n)) :=
|
||||
λk a, fiber_functor (@f k) (@f' k) (@p k) idp
|
||||
definition seq_diagram_over_fiber (g : Π⦃n⦄, A' n → A n)
|
||||
(p : Π⦃n⦄ (a : A' n), g (f' a) = f (g a)) : seq_diagram_over f (λn, fiber (@g n)) :=
|
||||
λk a, fiber_functor (@f' k) (@f k) (@p k) idp
|
||||
|
||||
definition seq_diagram_fiber (g : Π⦃n⦄, A n → A' n) (p : Π⦃n⦄ (a : A n), g (f a) = f' (g a))
|
||||
{n : ℕ} (a : A' n) : seq_diagram (λk, fiber (@g (n + k)) (rep f' k a)) :=
|
||||
definition seq_diagram_fiber (g : Π⦃n⦄, A' n → A n) (p : Π⦃n⦄ (a : A' n), g (f' a) = f (g a))
|
||||
{n : ℕ} (a : A n) : seq_diagram (λk, fiber (@g (n + k)) (rep f k a)) :=
|
||||
seq_diagram_of_over (seq_diagram_over_fiber g p) a
|
||||
|
||||
definition seq_diagram_fiber0 (g : Π⦃n⦄, A n → A' n) (p : Π⦃n⦄ (a : A n), g (f a) = f' (g a))
|
||||
(a : A' 0) : seq_diagram (λk, fiber (@g k) (rep0 f' k a)) :=
|
||||
λk, fiber_functor (@f k) (@f' k) (@p k) idp
|
||||
definition seq_diagram_fiber0 (g : Π⦃n⦄, A' n → A n) (p : Π⦃n⦄ (a : A' n), g (f' a) = f (g a))
|
||||
(a : A 0) : seq_diagram (λk, fiber (@g k) (rep0 f k a)) :=
|
||||
λk, fiber_functor (@f' k) (@f k) (@p k) idp
|
||||
|
||||
|
||||
end seq_colim
|
||||
|
|
|
@ -826,6 +826,32 @@ definition fiber_equiv_of_triangle {A B C : Type} {b : B} {f : A → B} {g : C
|
|||
(s : f ~ g ∘ h) : fiber f b ≃ fiber g b :=
|
||||
fiber_equiv_of_square h erfl s idp
|
||||
|
||||
definition is_trunc_fun_id (k : ℕ₋₂) (A : Type) : is_trunc_fun k (@id A) :=
|
||||
λa, is_trunc_of_is_contr _ _
|
||||
|
||||
definition is_conn_fun_id (k : ℕ₋₂) (A : Type) : is_conn_fun k (@id A) :=
|
||||
λa, _
|
||||
|
||||
open sigma.ops is_conn
|
||||
definition fiber_compose {A B C : Type} (g : B → C) (f : A → B) (c : C) :
|
||||
fiber (g ∘ f) c ≃ Σ(x : fiber g c), fiber f (point x) :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ intro x, exact ⟨fiber.mk (f (point x)) (point_eq x), fiber.mk (point x) idp⟩ },
|
||||
{ intro x, exact fiber.mk (point x.2) (ap g (point_eq x.2) ⬝ point_eq x.1) },
|
||||
{ intro x, induction x with x₁ x₂, induction x₁ with b p, induction x₂ with a q,
|
||||
induction p, esimp at q, induction q, reflexivity },
|
||||
{ intro x, induction x with a p, induction p, reflexivity }
|
||||
end
|
||||
|
||||
definition is_trunc_fun_compose (k : ℕ₋₂) {A B C : Type} {g : B → C} {f : A → B}
|
||||
(Hg : is_trunc_fun k g) (Hf : is_trunc_fun k f) : is_trunc_fun k (g ∘ f) :=
|
||||
λc, is_trunc_equiv_closed_rev k (fiber_compose g f c)
|
||||
|
||||
definition is_conn_fun_compose (k : ℕ₋₂) {A B C : Type} {g : B → C} {f : A → B}
|
||||
(Hg : is_conn_fun k g) (Hf : is_conn_fun k f) : is_conn_fun k (g ∘ f) :=
|
||||
λc, is_conn_equiv_closed_rev k (fiber_compose g f c) _
|
||||
|
||||
end fiber
|
||||
|
||||
namespace fin
|
||||
|
|
Loading…
Reference in a new issue