2016-01-22 14:53:34 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2016 Jakob von Raumer. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2016-01-23 19:15:59 +00:00
|
|
|
|
Authors: Jakob von Raumer, Ulrik Buchholtz
|
2016-01-22 14:53:34 +00:00
|
|
|
|
|
2016-02-15 21:05:31 +00:00
|
|
|
|
The Wedge Sum of Two pType Types
|
2016-01-22 14:53:34 +00:00
|
|
|
|
-/
|
2016-01-23 19:15:59 +00:00
|
|
|
|
import hit.pointed_pushout .connectedness
|
2016-01-22 14:53:34 +00:00
|
|
|
|
|
2016-03-02 22:19:44 +00:00
|
|
|
|
open eq pushout pointed unit trunc_index
|
2016-01-22 14:53:34 +00:00
|
|
|
|
|
2016-02-15 23:23:28 +00:00
|
|
|
|
definition pwedge (A B : Type*) : Type* := ppushout (pconst punit A) (pconst punit B)
|
2016-01-22 14:53:34 +00:00
|
|
|
|
|
2016-01-23 19:15:59 +00:00
|
|
|
|
namespace wedge
|
2016-01-25 16:54:48 +00:00
|
|
|
|
|
2016-01-22 14:53:34 +00:00
|
|
|
|
-- TODO maybe find a cleaner proof
|
2016-02-15 23:23:28 +00:00
|
|
|
|
protected definition unit (A : Type*) : A ≃* pwedge punit A :=
|
2016-01-22 14:53:34 +00:00
|
|
|
|
begin
|
2016-02-15 19:40:25 +00:00
|
|
|
|
fapply pequiv_of_pmap,
|
2016-01-22 14:53:34 +00:00
|
|
|
|
{ fapply pmap.mk, intro a, apply pinr a, apply respect_pt },
|
2016-02-15 19:40:25 +00:00
|
|
|
|
{ fapply is_equiv.adjointify, intro x, fapply pushout.elim_on x,
|
|
|
|
|
exact λ x, Point A, exact id, intro u, reflexivity,
|
|
|
|
|
intro x, fapply pushout.rec_on x, intro u, cases u, esimp, apply (glue unit.star)⁻¹,
|
2016-01-22 14:53:34 +00:00
|
|
|
|
intro a, reflexivity,
|
|
|
|
|
intro u, cases u, esimp, apply eq_pathover,
|
2016-02-15 19:40:25 +00:00
|
|
|
|
refine _ ⬝hp !ap_id⁻¹, fapply eq_hconcat, apply ap_compose inr,
|
|
|
|
|
krewrite elim_glue, fapply eq_hconcat, apply ap_idp, apply square_of_eq,
|
|
|
|
|
apply con.left_inv,
|
|
|
|
|
intro a, reflexivity},
|
2016-01-22 14:53:34 +00:00
|
|
|
|
end
|
2016-01-23 19:15:59 +00:00
|
|
|
|
end wedge
|
|
|
|
|
|
2016-02-08 11:07:53 +00:00
|
|
|
|
open trunc is_trunc is_conn function
|
|
|
|
|
|
2016-01-23 19:15:59 +00:00
|
|
|
|
namespace wedge_extension
|
|
|
|
|
section
|
|
|
|
|
-- The wedge connectivity lemma (Lemma 8.6.2)
|
2016-02-26 22:00:28 +00:00
|
|
|
|
parameters {A B : Type*} (n m : ℕ)
|
|
|
|
|
[cA : is_conn n A] [cB : is_conn m B]
|
2016-03-06 16:24:59 +00:00
|
|
|
|
(P : A → B → Type) [HP : Πa b, is_trunc (m + n) (P a b)]
|
2016-02-26 22:00:28 +00:00
|
|
|
|
(f : Πa : A, P a pt)
|
|
|
|
|
(g : Πb : B, P pt b)
|
|
|
|
|
(p : f pt = g pt)
|
2016-01-23 19:15:59 +00:00
|
|
|
|
|
2016-03-06 16:24:59 +00:00
|
|
|
|
include cA cB HP
|
|
|
|
|
private definition Q (a : A) : Type :=
|
|
|
|
|
fiber (λs : (Πb : B, P a b), s (Point B)) (f a)
|
2016-01-23 19:15:59 +00:00
|
|
|
|
|
2016-03-06 16:24:59 +00:00
|
|
|
|
private definition is_trunc_Q (a : A) : is_trunc (n.-1) (Q a) :=
|
|
|
|
|
begin
|
|
|
|
|
refine @is_conn.elim_general (m.-1) _ _ _ (P a) _ (f a),
|
|
|
|
|
rewrite [-succ_add_succ, of_nat_add_of_nat], intro b, apply HP
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
local attribute is_trunc_Q [instance]
|
2016-01-23 19:15:59 +00:00
|
|
|
|
private definition Q_sec : Πa : A, Q a :=
|
2016-03-06 16:24:59 +00:00
|
|
|
|
is_conn.elim (n.-1) Q (fiber.mk g p⁻¹)
|
2016-01-23 19:15:59 +00:00
|
|
|
|
|
|
|
|
|
protected definition ext : Π(a : A)(b : B), P a b :=
|
|
|
|
|
λa, fiber.point (Q_sec a)
|
2016-01-22 14:53:34 +00:00
|
|
|
|
|
2016-01-23 19:15:59 +00:00
|
|
|
|
protected definition β_left (a : A) : ext a (Point B) = f a :=
|
|
|
|
|
fiber.point_eq (Q_sec a)
|
|
|
|
|
|
|
|
|
|
private definition coh_aux : Σq : ext (Point A) = g,
|
|
|
|
|
β_left (Point A) = ap (λs : (Πb : B, P (Point A) b), s (Point B)) q ⬝ p⁻¹ :=
|
|
|
|
|
equiv.to_fun (fiber.fiber_eq_equiv (Q_sec (Point A)) (fiber.mk g p⁻¹))
|
2016-03-06 16:24:59 +00:00
|
|
|
|
(is_conn.elim_β (n.-1) Q (fiber.mk g p⁻¹))
|
2016-01-23 19:15:59 +00:00
|
|
|
|
|
|
|
|
|
protected definition β_right (b : B) : ext (Point A) b = g b :=
|
|
|
|
|
apd10 (sigma.pr1 coh_aux) b
|
|
|
|
|
|
|
|
|
|
private definition lem : β_left (Point A) = β_right (Point B) ⬝ p⁻¹ :=
|
|
|
|
|
begin
|
|
|
|
|
unfold β_right, unfold β_left,
|
|
|
|
|
krewrite (apd10_eq_ap_eval (sigma.pr1 coh_aux) (Point B)),
|
|
|
|
|
exact sigma.pr2 coh_aux,
|
2016-01-22 14:53:34 +00:00
|
|
|
|
end
|
2016-01-23 19:15:59 +00:00
|
|
|
|
|
|
|
|
|
protected definition coh
|
|
|
|
|
: (β_left (Point A))⁻¹ ⬝ β_right (Point B) = p :=
|
|
|
|
|
by rewrite [lem,con_inv,inv_inv,con.assoc,con.left_inv]
|
|
|
|
|
|
|
|
|
|
end
|
|
|
|
|
end wedge_extension
|