make everything compile on lean post 6f74f6522...

This commit is contained in:
Ulrik Buchholtz 2016-03-24 16:14:44 -04:00
parent 960e7075bd
commit 288e0d71b2
6 changed files with 21 additions and 74 deletions

View file

@ -1,5 +1,5 @@
import .LES_of_homotopy_groups homotopy.connectedness homotopy.homotopy_group homotopy.join
open eq is_trunc pointed homotopy is_equiv fiber equiv trunc nat chain_complex prod fin algebra
open eq is_trunc pointed is_conn is_equiv fiber equiv trunc nat chain_complex prod fin algebra
group trunc_index function join pushout
namespace nat
@ -87,58 +87,23 @@ namespace is_conn
definition join_empty_right [constructor] (A : Type) : join A empty ≃ A :=
begin
fapply equiv.MK,
{ intro x, induction x with a o v,
{ exact a},
{ exact empty.elim o},
{ exact empty.elim (pr2 v)}},
{ exact pushout.inl},
{ intro x, induction x with a o a o,
{ exact a },
{ exact empty.elim o },
{ exact empty.elim o } },
{ exact pushout.inl },
{ intro a, reflexivity},
{ intro x, induction x with a o v,
{ reflexivity},
{ exact empty.elim o},
{ exact empty.elim (pr2 v)}}
{ intro x, induction x with a o a o,
{ reflexivity },
{ exact empty.elim o },
{ exact empty.elim o } }
end
definition join_functor [unfold 7] {A A' B B' : Type} (f : A → A') (g : B → B') :
join A B → join A' B' :=
begin
intro x, induction x with a b v,
{ exact inl (f a)},
{ exact inr (g b)},
{ exact glue (f (pr1 v), g (pr2 v))}
end
theorem join_functor_glue {A A' B B' : Type} (f : A → A') (g : B → B')
(v : A × B) : ap (join_functor f g) (glue v) = glue (f (pr1 v), g (pr2 v)) :=
!elim_glue
definition natural_square2 {A B X : Type} {f : A → X} {g : B → X} (h : Πa b, f a = g b)
{a a' : A} {b b' : B} (p : a = a') (q : b = b')
: square (ap f p) (ap g q) (h a b) (h a' b') :=
by induction p; induction q; exact hrfl
definition join_equiv_join {A A' B B' : Type} (f : A ≃ A') (g : B ≃ B') :
join A B ≃ join A' B' :=
begin
fapply equiv.MK,
{ apply join_functor f g},
{ apply join_functor f⁻¹ g⁻¹},
{ intro x', induction x' with a' b' v',
{ esimp, exact ap inl (right_inv f a')},
{ esimp, exact ap inr (right_inv g b')},
{ cases v' with a' b', apply eq_pathover,
rewrite [▸*, ap_id, ap_compose' (join_functor _ _), ▸*],
xrewrite [+join_functor_glue, ▸*],
exact natural_square2 jglue (right_inv f a') (right_inv g b')}},
{ intro x, induction x with a b v,
{ esimp, exact ap inl (left_inv f a)},
{ esimp, exact ap inr (left_inv g b)},
{ cases v with a b, apply eq_pathover,
rewrite [▸*, ap_id, ap_compose' (join_functor _ _), ▸*],
xrewrite [+join_functor_glue, ▸*],
exact natural_square2 jglue (left_inv f a) (left_inv g b)}}
end
section
open sphere sphere_index
@ -149,19 +114,6 @@ namespace is_conn
definition succ_add_plus_one (n m : ℕ₋₁) : (n.+1) +1+ m = (n +1+ m).+1 :=
begin induction m with m IH, reflexivity, exact ap succ IH end
definition sphere_join_sphere (n m : ℕ₋₁) : join (sphere n) (sphere m) ≃ sphere (n +1+ m) :=
begin
revert n, induction m with m IH: intro n,
{ apply join_empty_right},
{ rewrite [sphere_succ m],
refine join_equiv_join erfl !join.bool⁻¹ᵉ ⬝e _,
refine !join.assoc⁻¹ᵉ ⬝e _,
refine join_equiv_join !join.symm erfl ⬝e _,
refine join_equiv_join !join.bool erfl ⬝e _,
rewrite [-sphere_succ n],
refine !IH ⬝e _,
rewrite [add_plus_one_succ, succ_add_plus_one]}
end
end
end is_conn

View file

@ -208,9 +208,10 @@ namespace chain_complex
theorem fiber_sequence_fun_eq : Π(x : fiber_sequence_carrier f (n + 4)),
fiber_sequence_carrier_pequiv f n (fiber_sequence_fun f (n + 3) x) =
ap1 (fiber_sequence_fun f n) (fiber_sequence_carrier_pequiv f (n + 1) x)⁻¹ :=
homotopy_of_inv_homotopy
(pequiv.to_equiv (fiber_sequence_carrier_pequiv f (n + 1)))
(fiber_sequence_fun_eq_helper f n)
begin
apply homotopy_of_inv_homotopy_pre (fiber_sequence_carrier_pequiv f (n + 1)),
apply fiber_sequence_fun_eq_helper f n
end
theorem fiber_sequence_fun_phomotopy :
fiber_sequence_carrier_pequiv f n ∘*

View file

@ -209,8 +209,8 @@ namespace chain_complex
have H2 : tcc_to_fn X (f m) ((equiv_of_eq (ap (λx, X x) (c m)))⁻¹ᵉ (e⁻¹ y)) = pt,
begin
refine _ ⬝ ap e⁻¹ᵉ* q ⬝ (respect_pt (e⁻¹ᵉ*)), apply eq_inv_of_eq, clear q, revert y,
refine inv_homotopy_of_homotopy (pequiv.to_equiv e) _,
apply inv_homotopy_of_homotopy, apply p
apply inv_homotopy_of_homotopy_pre e,
apply inv_homotopy_of_homotopy_pre, apply p
end,
induction (H _ H2) with x r,
refine fiber.mk (e (cast (ap (λx, X x) (c (S m))) (cast (ap (λx, X (S x)) (c m)) x))) _,

View file

@ -2,7 +2,7 @@
import homotopy.connectedness homotopy.join
open eq sigma pi function join homotopy is_trunc equiv is_equiv
open eq sigma pi function join is_conn is_trunc equiv is_equiv
namespace retraction
variables {A B C : Type} (r2 : B → C) (r1 : A → B)

View file

@ -2,7 +2,7 @@
import homotopy.wedge types.pi .LES_applications --TODO: remove
open eq homotopy is_trunc pointed susp nat pi equiv is_equiv trunc fiber trunc_index
open eq is_conn is_trunc pointed susp nat pi equiv is_equiv trunc fiber trunc_index
-- definition iterated_loop_ptrunc_pequiv_con' (n : ℕ₋₂) (k : ) (A : Type*)
-- (p q : Ω[k](ptrunc (n+k) (Ω A))) :
@ -253,8 +253,8 @@ namespace sphere
exact !phomotopy_group_pequiv_loop_ptrunc⁻¹ᵉ*,
end
print phomotopy_group_pequiv_loop_ptrunc
print iterated_loop_ptrunc_pequiv
--print phomotopy_group_pequiv_loop_ptrunc
--print iterated_loop_ptrunc_pequiv
-- definition to_fun_stability_pequiv (k n : ) (H : k + 3 ≤ 2 * n) --(p : π*[k + 1] (S. (n+1)))
-- : stability_pequiv (k+1) n H = _ ∘ _ ∘ cast (ap (ptrunc 0) (loop_space_succ_eq_in (S. (n+1)) (k+1))) :=
-- sorry

View file

@ -144,13 +144,7 @@ namespace spherical_fibrations
definition tau : S 2 → BG 2 :=
begin
intro v, induction v with x, do 2 exact pt,
fapply sigma_eq,
{ apply ua, fapply equiv.MK,
{ },
{ },
{ },
{ } },
{ }
exact sorry
end
end two_sphere