move more results and make arguments explicit

More results from the Spectral repository are moved to this library
Also make various type-class arguments of truncatedness and equivalences which were hard to synthesize explicit
This commit is contained in:
Floris van Doorn 2018-09-11 16:45:30 +02:00
parent 14c8fbfea3
commit 9a17a244c9
39 changed files with 1179 additions and 520 deletions

View file

@ -21,6 +21,7 @@ The following files are [ported](../port.md) from the standard library. If anyth
Files which are not ported from the standard library: Files which are not ported from the standard library:
* [inf_group](inf_group.hlean) : algebraic structures which are not assumes to be sets. No higher coherences are assumed. Truncated algebraic structures extend these structures with the assumption that they are sets. * [inf_group](inf_group.hlean) : algebraic structures which are not assumes to be sets. No higher coherences are assumed. Truncated algebraic structures extend these structures with the assumption that they are sets.
* [inf_group_theory](inf_group_theory.hlean) : Some very basic group theory using InfGroups
* [group_theory](group_theory.hlean) : Basic theorems about group homomorphisms and isomorphisms * [group_theory](group_theory.hlean) : Basic theorems about group homomorphisms and isomorphisms
* [trunc_group](trunc_group.hlean) : truncate an infinity-group to a group * [trunc_group](trunc_group.hlean) : truncate an infinity-group to a group
* [homotopy_group](homotopy_group.hlean) : homotopy groups of a pointed type * [homotopy_group](homotopy_group.hlean) : homotopy groups of a pointed type

View file

@ -232,3 +232,9 @@ attribute Ring.carrier [coercion]
attribute Ring.struct [instance] attribute Ring.struct [instance]
end algebra end algebra
open algebra
namespace infgroup
attribute [coercion] InfGroup_of_Group
attribute [coercion] AbInfGroup_of_AbGroup
end infgroup

View file

@ -78,9 +78,7 @@ namespace category
definition is_trunc_1_ob : is_trunc 1 ob := definition is_trunc_1_ob : is_trunc 1 ob :=
begin begin
apply is_trunc_succ_intro, intro a b, apply is_trunc_succ_intro, intro a b,
fapply is_trunc_is_equiv_closed, exact is_trunc_equiv_closed_rev _ (eq_equiv_iso a b) _
exact (@eq_of_iso _ _ a b),
apply is_equiv_inv,
end end
end basic end basic

View file

@ -68,14 +68,14 @@ namespace category
definition is_univalent_Set (A B : set) : is_equiv (iso_of_eq : A = B → A ≅ B) := definition is_univalent_Set (A B : set) : is_equiv (iso_of_eq : A = B → A ≅ B) :=
have H₁ : is_equiv (@iso_of_equiv A B ∘ @equiv_of_eq A B ∘ subtype_eq_inv _ _ ∘ have H₁ : is_equiv (@iso_of_equiv A B ∘ @equiv_of_eq A B ∘ subtype_eq_inv _ _ ∘
@ap _ _ (to_fun (trunctype.sigma_char 0)) A B), from @ap _ _ (to_fun (trunctype.sigma_char 0)) A B), from
@is_equiv_compose _ _ _ _ _ is_equiv_compose _ _
(@is_equiv_compose _ _ _ _ _ (is_equiv_compose _ _
(@is_equiv_compose _ _ _ _ _ (is_equiv_compose _ _
_ _
(@is_equiv_subtype_eq_inv _ _ _ _ _)) (@is_equiv_subtype_eq_inv _ _ _ _ _))
!univalence) !univalence)
!is_equiv_iso_of_equiv, !is_equiv_iso_of_equiv,
is_equiv.homotopy_closed _ (iso_of_eq_eq_compose A B)⁻¹ʰᵗʸ is_equiv.homotopy_closed _ (iso_of_eq_eq_compose A B)⁻¹ʰᵗʸ _
end set end set

View file

@ -143,7 +143,7 @@ namespace category
definition fully_faithful_equiv (F : C ⇒ D) : fully_faithful F ≃ (faithful F × full F) := definition fully_faithful_equiv (F : C ⇒ D) : fully_faithful F ≃ (faithful F × full F) :=
equiv_of_is_prop (λH, (faithful_of_fully_faithful F, full_of_fully_faithful F)) equiv_of_is_prop (λH, (faithful_of_fully_faithful F, full_of_fully_faithful F))
(λH, fully_faithful_of_full_of_faithful (pr1 H) (pr2 H)) (λH, fully_faithful_of_full_of_faithful (pr1 H) (pr2 H)) _ _
/- alternative proof using direct calculation with equivalences /- alternative proof using direct calculation with equivalences
@ -165,7 +165,7 @@ namespace category
definition fully_faithful_compose (G : D ⇒ E) (F : C ⇒ D) [fully_faithful G] [fully_faithful F] : definition fully_faithful_compose (G : D ⇒ E) (F : C ⇒ D) [fully_faithful G] [fully_faithful F] :
fully_faithful (G ∘f F) := fully_faithful (G ∘f F) :=
λc c', is_equiv_compose (to_fun_hom G) (to_fun_hom F) λc c', is_equiv_compose (to_fun_hom G) (to_fun_hom F) _ _
definition full_compose (G : D ⇒ E) (F : C ⇒ D) [full G] [full F] : full (G ∘f F) := definition full_compose (G : D ⇒ E) (F : C ⇒ D) [full G] [full F] : full (G ∘f F) :=
λc c', is_surjective_compose (to_fun_hom G) (to_fun_hom F) _ _ λc c', is_surjective_compose (to_fun_hom G) (to_fun_hom F) _ _

View file

@ -201,10 +201,7 @@ namespace iso
-- The type of isomorphisms between two objects is a set -- The type of isomorphisms between two objects is a set
definition is_set_iso [instance] : is_set (a ≅ b) := definition is_set_iso [instance] : is_set (a ≅ b) :=
begin is_trunc_equiv_closed _ !iso.sigma_char _
apply is_trunc_is_equiv_closed,
apply equiv.to_is_equiv (!iso.sigma_char),
end
definition iso_of_eq [unfold 5] (p : a = b) : a ≅ b := definition iso_of_eq [unfold 5] (p : a = b) : a ≅ b :=
eq.rec_on p (iso.refl a) eq.rec_on p (iso.refl a)

View file

@ -93,7 +93,7 @@ namespace nat_trans
end end
definition is_set_nat_trans [instance] : is_set (F ⟹ G) := definition is_set_nat_trans [instance] : is_set (F ⟹ G) :=
by apply is_trunc_is_equiv_closed; apply (equiv.to_is_equiv !nat_trans.sigma_char) is_trunc_equiv_closed _ !nat_trans.sigma_char _
definition change_natural_map [constructor] (η : F ⟹ G) (f : Π (a : C), F a ⟶ G a) definition change_natural_map [constructor] (η : F ⟹ G) (f : Π (a : C), F a ⟶ G a)
(p : Πa, η a = f a) : F ⟹ G := (p : Πa, η a = f a) : F ⟹ G :=

View file

@ -1,7 +1,7 @@
/- /-
Copyright (c) 2014 Jeremy Avigad. All rights reserved. Copyright (c) 2014 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn
Various multiplicative and additive structures. Partially modeled on Isabelle's library. Various multiplicative and additive structures. Partially modeled on Isabelle's library.
-/ -/

View file

@ -6,9 +6,10 @@ Authors: Floris van Doorn
Basic group theory Basic group theory
-/ -/
import algebra.category.category algebra.bundled .homomorphism types.pointed2 import algebra.category.category algebra.inf_group_theory .homomorphism types.pointed2
algebra.trunc_group
open eq algebra pointed function is_trunc pi equiv is_equiv open eq algebra pointed function is_trunc pi equiv is_equiv sigma sigma.ops trunc
set_option class.force_new true set_option class.force_new true
namespace group namespace group
@ -155,6 +156,10 @@ namespace group
infixr ` ∘g `:75 := homomorphism_compose infixr ` ∘g `:75 := homomorphism_compose
notation 1 := homomorphism_id _ notation 1 := homomorphism_id _
definition homomorphism_compose_eq (ψ : G₂ →g G₃) (φ : G₁ →g G₂) (g : G₁) :
(ψ ∘g φ) g = ψ (φ g) :=
by reflexivity
structure isomorphism (A B : Group) := structure isomorphism (A B : Group) :=
(to_hom : A →g B) (to_hom : A →g B)
(is_equiv_to_hom : is_equiv to_hom) (is_equiv_to_hom : is_equiv to_hom)
@ -175,9 +180,9 @@ namespace group
(p : Πg₁ g₂, φ (g₁ * g₂) = φ g₁ * φ g₂) : G₁ ≃g G₂ := (p : Πg₁ g₂, φ (g₁ * g₂) = φ g₁ * φ g₂) : G₁ ≃g G₂ :=
isomorphism.mk (homomorphism.mk φ p) !to_is_equiv isomorphism.mk (homomorphism.mk φ p) !to_is_equiv
definition isomorphism_of_eq [constructor] {G₁ G₂ : Group} (φ : G₁ = G₂) : G₁ ≃g G₂ := definition isomorphism.MK [constructor] (φ : G₁ →g G₂) (ψ : G₂ →g G₁)
isomorphism_of_equiv (equiv_of_eq (ap Group.carrier φ)) (p : φ ∘g ψ ~ gid G₂) (q : ψ ∘g φ ~ gid G₁) : G₁ ≃g G₂ :=
begin intros, induction φ, reflexivity end isomorphism.mk φ (adjointify φ ψ p q)
definition to_ginv [constructor] (φ : G₁ ≃g G₂) : G₂ →g G₁ := definition to_ginv [constructor] (φ : G₁ ≃g G₂) : G₂ →g G₁ :=
homomorphism.mk φ⁻¹ homomorphism.mk φ⁻¹
@ -186,6 +191,13 @@ namespace group
rewrite [respect_mul φ, +right_inv φ] rewrite [respect_mul φ, +right_inv φ]
end end end end
definition isomorphism_of_eq [constructor] {G₁ G₂ : Group} (φ : G₁ = G₂) : G₁ ≃g G₂ :=
isomorphism_of_equiv (equiv_of_eq (ap Group.carrier φ))
begin intros, induction φ, reflexivity end
definition isomorphism_ap {A : Type} (F : A → Group) {a b : A} (p : a = b) : F a ≃g F b :=
isomorphism_of_eq (ap F p)
variable (G) variable (G)
definition isomorphism.refl [refl] [constructor] : G ≃g G := definition isomorphism.refl [refl] [constructor] : G ≃g G :=
isomorphism.mk 1 !is_equiv_id isomorphism.mk 1 !is_equiv_id
@ -195,7 +207,7 @@ namespace group
isomorphism.mk (to_ginv φ) !is_equiv_inv isomorphism.mk (to_ginv φ) !is_equiv_inv
definition isomorphism.trans [trans] [constructor] (φ : G₁ ≃g G₂) (ψ : G₂ ≃g G₃) : G₁ ≃g G₃ := definition isomorphism.trans [trans] [constructor] (φ : G₁ ≃g G₂) (ψ : G₂ ≃g G₃) : G₁ ≃g G₃ :=
isomorphism.mk (ψ ∘g φ) !is_equiv_compose isomorphism.mk (ψ ∘g φ) (is_equiv_compose ψ φ _ _)
definition isomorphism.eq_trans [trans] [constructor] definition isomorphism.eq_trans [trans] [constructor]
{G₁ G₂ : Group} {G₃ : Group} (φ : G₁ = G₂) (ψ : G₂ ≃g G₃) : G₁ ≃g G₃ := {G₁ G₂ : Group} {G₃ : Group} (φ : G₁ = G₂) (ψ : G₂ ≃g G₃) : G₁ ≃g G₃ :=
@ -271,6 +283,45 @@ namespace group
apply phomotopy_of_homotopy, reflexivity apply phomotopy_of_homotopy, reflexivity
end end
protected definition homomorphism.sigma_char [constructor]
(A B : Group) : (A →g B) ≃ Σ(f : A → B), is_mul_hom f :=
begin
fapply equiv.MK,
{intro F, exact ⟨F, _⟩ },
{intro p, cases p with f H, exact (homomorphism.mk f H) },
{intro p, cases p, reflexivity },
{intro F, cases F, reflexivity },
end
definition homomorphism_pathover {A : Type} {a a' : A} (p : a = a')
{B : A → Group} {C : A → Group} (f : B a →g C a) (g : B a' →g C a')
(r : homomorphism.φ f =[p] homomorphism.φ g) : f =[p] g :=
begin
fapply pathover_of_fn_pathover_fn,
{ intro a, apply homomorphism.sigma_char },
{ fapply sigma_pathover, exact r, apply is_prop.elimo }
end
protected definition isomorphism.sigma_char [constructor]
(A B : Group) : (A ≃g B) ≃ Σ(f : A →g B), is_equiv f :=
begin
fapply equiv.MK,
{intro F, exact ⟨F, _⟩ },
{intro p, exact (isomorphism.mk p.1 p.2) },
{intro p, cases p, reflexivity },
{intro F, cases F, reflexivity },
end
definition isomorphism_pathover {A : Type} {a a' : A} (p : a = a')
{B : A → Group} {C : A → Group} (f : B a ≃g C a) (g : B a' ≃g C a')
(r : pathover (λa, B a → C a) f p g) : f =[p] g :=
begin
fapply pathover_of_fn_pathover_fn,
{ intro a, apply isomorphism.sigma_char },
{ fapply sigma_pathover, apply homomorphism_pathover, exact r, apply is_prop.elimo }
end
definition isomorphism_eq {G H : Group} {φ ψ : G ≃g H} (p : φ ~ ψ) : φ = ψ := definition isomorphism_eq {G H : Group} {φ ψ : G ≃g H} (p : φ ~ ψ) : φ = ψ :=
begin begin
induction φ with φ φe, induction ψ with ψ ψe, induction φ with φ φe, induction ψ with ψ ψe,
@ -329,6 +380,30 @@ namespace group
definition aghomomorphism [constructor] (G H : AbGroup) : AbGroup := definition aghomomorphism [constructor] (G H : AbGroup) : AbGroup :=
AbGroup.mk (G →g H) (ab_group_homomorphism G H) AbGroup.mk (G →g H) (ab_group_homomorphism G H)
infixr ` →gg `:56 := aghomomorphism
/- some properties of binary homomorphisms -/
definition pmap_of_homomorphism2 [constructor] {G H K : AbGroup} (φ : G →g H →gg K) :
G →* H →** K :=
pmap.mk (λg, pmap_of_homomorphism (φ g))
(eq_of_phomotopy (phomotopy_of_homotopy (ap010 group_fun (to_respect_one φ))))
definition homomorphism_apply [constructor] (G H : AbGroup) (g : G) :
(G →gg H) →g H :=
begin
fapply homomorphism.mk,
{ intro φ, exact φ g },
{ intros φ φ', reflexivity }
end
definition homomorphism_swap [constructor] {G H K : AbGroup} (φ : G →g H →gg K) :
H →g G →gg K :=
begin
fapply homomorphism.mk,
{ intro h, exact homomorphism_apply H K h ∘g φ },
{ intro h h', apply homomorphism_eq, intro g, exact to_respect_mul (φ g) h h' }
end
/- given an equivalence A ≃ B we can transport a group structure on A to a group structure on B -/ /- given an equivalence A ≃ B we can transport a group structure on A to a group structure on B -/
section section
@ -409,22 +484,25 @@ namespace group
fapply AbGroup_of_Group trivial_group, intro x y, reflexivity fapply AbGroup_of_Group trivial_group, intro x y, reflexivity
end end
definition trivial_group_of_is_contr [H : is_contr G] : G ≃g G0 := definition trivial_group_of_is_contr (H : is_contr G) : G ≃g G0 :=
begin begin
fapply isomorphism_of_equiv, fapply isomorphism_of_equiv,
{ apply equiv_unit_of_is_contr}, { exact equiv_unit_of_is_contr _ _ },
{ intros, reflexivity} { intros, reflexivity }
end end
definition ab_group_of_is_contr (A : Type) [is_contr A] : ab_group A := definition isomorphism_of_is_contr {G H : Group} (hG : is_contr G) (hH : is_contr H) : G ≃g H :=
have ab_group unit, from ab_group_unit, trivial_group_of_is_contr G _ ⬝g (trivial_group_of_is_contr H _)⁻¹ᵍ
ab_group_equiv_closed (equiv_unit_of_is_contr A)⁻¹ᵉ
definition group_of_is_contr (A : Type) [is_contr A] : group A := definition ab_group_of_is_contr (A : Type) (H : is_contr A) : ab_group A :=
have ab_group A, from ab_group_of_is_contr A, by apply _ have ab_group unit, from ab_group_unit,
ab_group_equiv_closed (equiv_unit_of_is_contr A _)⁻¹ᵉ
definition group_of_is_contr (A : Type) (H : is_contr A) : group A :=
have ab_group A, from ab_group_of_is_contr A H, by apply _
definition ab_group_lift_unit : ab_group (lift unit) := definition ab_group_lift_unit : ab_group (lift unit) :=
ab_group_of_is_contr (lift unit) ab_group_of_is_contr (lift unit) _
definition trivial_ab_group_lift : AbGroup := definition trivial_ab_group_lift : AbGroup :=
AbGroup.mk _ ab_group_lift_unit AbGroup.mk _ ab_group_lift_unit
@ -472,6 +550,9 @@ namespace group
mul_left_inv_pt := mul.left_inv⦄ mul_left_inv_pt := mul.left_inv⦄
end end
definition pgroup_of_Group (X : Group) : pgroup X :=
pgroup_of_group _ idp
definition Group_of_pgroup (G : Type*) [pgroup G] : Group := definition Group_of_pgroup (G : Type*) [pgroup G] : Group :=
Group.mk G _ Group.mk G _
@ -481,39 +562,6 @@ namespace group
mul_pt := mul_one, mul_pt := mul_one,
mul_left_inv_pt := mul.left_inv ⦄ mul_left_inv_pt := mul.left_inv ⦄
-- infinity pgroups
structure inf_pgroup [class] (X : Type*) extends inf_semigroup X, has_inv X :=
(pt_mul : Πa, mul pt a = a)
(mul_pt : Πa, mul a pt = a)
(mul_left_inv_pt : Πa, mul (inv a) a = pt)
definition inf_group_of_inf_pgroup [reducible] [instance] (X : Type*) [H : inf_pgroup X]
: inf_group X :=
⦃inf_group, H,
one := pt,
one_mul := inf_pgroup.pt_mul ,
mul_one := inf_pgroup.mul_pt,
mul_left_inv := inf_pgroup.mul_left_inv_pt⦄
definition inf_pgroup_of_inf_group (X : Type*) [H : inf_group X] (p : one = pt :> X) : inf_pgroup X :=
begin
cases X with X x, esimp at *, induction p,
exact ⦃inf_pgroup, H,
pt_mul := one_mul,
mul_pt := mul_one,
mul_left_inv_pt := mul.left_inv⦄
end
definition inf_Group_of_inf_pgroup (G : Type*) [inf_pgroup G] : InfGroup :=
InfGroup.mk G _
definition inf_pgroup_InfGroup [instance] (G : InfGroup) : inf_pgroup G :=
⦃ inf_pgroup, InfGroup.struct G,
pt_mul := one_mul,
mul_pt := mul_one,
mul_left_inv_pt := mul.left_inv ⦄
/- equality of groups and abelian groups -/ /- equality of groups and abelian groups -/
definition group.to_has_mul {A : Type} (H : group A) : has_mul A := _ definition group.to_has_mul {A : Type} (H : group A) : has_mul A := _
@ -610,7 +658,7 @@ namespace group
end end
definition trivial_group_of_is_contr' (G : Group) [H : is_contr G] : G = G0 := definition trivial_group_of_is_contr' (G : Group) [H : is_contr G] : G = G0 :=
eq_of_isomorphism (trivial_group_of_is_contr G) eq_of_isomorphism (trivial_group_of_is_contr G _)
definition pequiv_of_isomorphism_of_eq {G₁ G₂ : Group} (p : G₁ = G₂) : definition pequiv_of_isomorphism_of_eq {G₁ G₂ : Group} (p : G₁ = G₂) :
pequiv_of_isomorphism (isomorphism_of_eq p) = pequiv_of_eq (ap pType_of_Group p) := pequiv_of_isomorphism (isomorphism_of_eq p) = pequiv_of_eq (ap pType_of_Group p) :=
@ -622,4 +670,38 @@ namespace group
{ apply is_prop.elim } { apply is_prop.elim }
end end
/- relation with infgroups -/
-- todo: define homomorphism in terms of inf_homomorphism and similar for isomorphism?
open infgroup
definition homomorphism_of_inf_homomorphism [constructor] {G H : Group} (φ : G →∞g H) : G →g H :=
homomorphism.mk φ (inf_homomorphism.struct φ)
definition inf_homomorphism_of_homomorphism [constructor] {G H : Group} (φ : G →g H) : G →∞g H :=
inf_homomorphism.mk φ (homomorphism.struct φ)
definition isomorphism_of_inf_isomorphism [constructor] {G H : Group} (φ : G ≃∞g H) : G ≃g H :=
isomorphism.mk (homomorphism_of_inf_homomorphism φ) (inf_isomorphism.is_equiv_to_hom φ)
definition inf_isomorphism_of_isomorphism [constructor] {G H : Group} (φ : G ≃g H) : G ≃∞g H :=
inf_isomorphism.mk (inf_homomorphism_of_homomorphism φ) (isomorphism.is_equiv_to_hom φ)
definition gtrunc_functor {A B : InfGroup} (f : A →∞g B) : gtrunc A →g gtrunc B :=
begin
apply homomorphism.mk (trunc_functor 0 f),
intros x x', induction x with a, induction x' with a', apply ap tr, exact respect_mul f a a'
end
definition gtrunc_isomorphism_gtrunc {A B : InfGroup} (f : A ≃∞g B) : gtrunc A ≃g gtrunc B :=
isomorphism_of_equiv (trunc_equiv_trunc 0 (equiv_of_inf_isomorphism f))
(to_respect_mul (gtrunc_functor f))
definition gtr [constructor] (X : InfGroup) : X →∞g gtrunc X :=
inf_homomorphism.mk tr homotopy2.rfl
definition gtrunc_isomorphism [constructor] (X : InfGroup) [H : is_set X] : gtrunc X ≃∞g X :=
(inf_isomorphism_of_equiv (trunc_equiv 0 X)⁻¹ᵉ homotopy2.rfl)⁻¹ᵍ⁸
definition is_set_group_inf [instance] (G : Group) : group G := Group.struct G
end group end group

View file

@ -49,8 +49,7 @@ section add_group_A_B
by rewrite [*sub_eq_add_neg, *(respect_add f), (respect_neg f)] by rewrite [*sub_eq_add_neg, *(respect_add f), (respect_neg f)]
definition is_embedding_of_is_add_hom [add_group B] (f : A → B) [is_add_hom f] definition is_embedding_of_is_add_hom [add_group B] (f : A → B) [is_add_hom f]
(H : ∀ x, f x = 0 → x = 0) : (H : ∀ x, f x = 0 → x = 0) : is_embedding f :=
is_embedding f :=
is_embedding_of_is_injective is_embedding_of_is_injective
(take x₁ x₂, (take x₁ x₂,
suppose f x₁ = f x₂, suppose f x₁ = f x₂,

View file

@ -10,25 +10,10 @@ import .trunc_group types.trunc .group_theory types.nat.hott
open nat eq pointed trunc is_trunc algebra group function equiv unit is_equiv nat open nat eq pointed trunc is_trunc algebra group function equiv unit is_equiv nat
/- todo: prove more properties of homotopy groups using gtrunc and agtrunc -/
namespace eq namespace eq
definition inf_pgroup_loop [constructor] [instance] (A : Type*) : inf_pgroup (Ω A) :=
inf_pgroup.mk concat con.assoc inverse idp_con con_idp con.left_inv
definition inf_group_loop [constructor] (A : Type*) : inf_group (Ω A) := _
definition ab_inf_group_loop [constructor] [instance] (A : Type*) : ab_inf_group (Ω (Ω A)) :=
⦃ab_inf_group, inf_group_loop _, mul_comm := eckmann_hilton⦄
definition inf_group_loopn (n : ) (A : Type*) [H : is_succ n] : inf_group (Ω[n] A) :=
by induction H; exact _
definition ab_inf_group_loopn (n : ) (A : Type*) [H : is_at_least_two n] : ab_inf_group (Ω[n] A) :=
by induction H; exact _
definition gloop [constructor] (A : Type*) : InfGroup :=
InfGroup.mk (Ω A) (inf_group_loop A)
definition homotopy_group [reducible] [constructor] (n : ) (A : Type*) : Set* := definition homotopy_group [reducible] [constructor] (n : ) (A : Type*) : Set* :=
ptrunc 0 (Ω[n] A) ptrunc 0 (Ω[n] A)
@ -36,9 +21,9 @@ namespace eq
section section
local attribute inf_group_loopn [instance] local attribute inf_group_loopn [instance]
definition group_homotopy_group [instance] [constructor] [reducible] (n : ) [is_succ n] (A : Type*) definition group_homotopy_group [instance] [constructor] [reducible] (n : ) [is_succ n]
: group (π[n] A) := (A : Type*) : group (π[n] A) :=
trunc_group (Ω[n] A) group_trunc (Ω[n] A)
end end
definition group_homotopy_group2 [instance] (k : ) (A : Type*) : definition group_homotopy_group2 [instance] (k : ) (A : Type*) :
@ -47,26 +32,25 @@ namespace eq
section section
local attribute ab_inf_group_loopn [instance] local attribute ab_inf_group_loopn [instance]
definition ab_group_homotopy_group [constructor] [reducible] (n : ) [is_at_least_two n] (A : Type*) definition ab_group_homotopy_group [constructor] [reducible] (n : ) [is_at_least_two n]
: ab_group (π[n] A) := (A : Type*) : ab_group (π[n] A) :=
trunc_ab_group (Ω[n] A) ab_group_trunc (Ω[n] A)
end end
local attribute ab_group_homotopy_group [instance] local attribute ab_group_homotopy_group [instance]
definition ghomotopy_group [constructor] (n : ) [is_succ n] (A : Type*) : Group := definition ghomotopy_group [constructor] (n : ) [is_succ n] (A : Type*) : Group :=
Group.mk (π[n] A) _ gtrunc (Ωg[n] A)
definition aghomotopy_group [constructor] (n : ) [is_at_least_two n] (A : Type*) : AbGroup := definition aghomotopy_group [constructor] (n : ) [is_at_least_two n] (A : Type*) : AbGroup :=
AbGroup.mk (π[n] A) _ agtrunc (Ωag[n] A)
definition fundamental_group [constructor] (A : Type*) : Group :=
ghomotopy_group 1 A
notation `πg[`:95 n:0 `]`:0 := ghomotopy_group n notation `πg[`:95 n:0 `]`:0 := ghomotopy_group n
notation `πag[`:95 n:0 `]`:0 := aghomotopy_group n notation `πag[`:95 n:0 `]`:0 := aghomotopy_group n
notation `π₁` := fundamental_group -- should this be notation for the group or pointed type? definition fundamental_group [constructor] (A : Type*) : Group := πg[1] A
notation `π₁` := fundamental_group
definition tr_mul_tr {n : } {A : Type*} (p q : Ω[n + 1] A) : definition tr_mul_tr {n : } {A : Type*} (p q : Ω[n + 1] A) :
tr p *[πg[n+1] A] tr q = tr (p ⬝ q) := tr p *[πg[n+1] A] tr q = tr (p ⬝ q) :=
@ -105,8 +89,8 @@ namespace eq
begin begin
apply trivial_group_of_is_contr, apply trivial_group_of_is_contr,
apply is_trunc_trunc_of_is_trunc, apply is_trunc_trunc_of_is_trunc,
apply is_contr_loop_of_is_trunc, apply is_contr_loop_of_is_trunc (n+1),
apply is_trunc_succ_succ_of_is_set exact is_trunc_succ_succ_of_is_set _ _ _
end end
definition homotopy_group_succ_out (n : ) (A : Type*) : π[n + 1] A = π₁ (Ω[n] A) := idp definition homotopy_group_succ_out (n : ) (A : Type*) : π[n + 1] A = π₁ (Ω[n] A) := idp
@ -136,7 +120,7 @@ namespace eq
begin begin
apply is_trunc_trunc_of_is_trunc, apply is_trunc_trunc_of_is_trunc,
apply is_contr_loop_of_is_trunc, apply is_contr_loop_of_is_trunc,
apply is_trunc_of_is_contr exact is_trunc_of_is_contr _ _ _
end end
definition homotopy_group_functor [constructor] (n : ) {A B : Type*} (f : A →* B) definition homotopy_group_functor [constructor] (n : ) {A B : Type*} (f : A →* B)
@ -192,6 +176,11 @@ namespace eq
{ reflexivity} { reflexivity}
end end
/- maybe rename: ghomotopy_group_functor -/
definition homotopy_group_homomorphism [constructor] (n : ) [H : is_succ n] {A B : Type*}
(f : A →* B) : πg[n] A →g πg[n] B :=
gtrunc_functor (Ωg→[n] f)
definition homotopy_group_functor_mul [constructor] (n : ) {A B : Type*} (g : A →* B) definition homotopy_group_functor_mul [constructor] (n : ) {A B : Type*} (g : A →* B)
(p q : πg[n+1] A) : (p q : πg[n+1] A) :
(π→[n + 1] g) (p *[πg[n+1] A] q) = (π→[n+1] g) p *[πg[n+1] B] (π→[n + 1] g) q := (π→[n + 1] g) (p *[πg[n+1] A] q) = (π→[n+1] g) p *[πg[n+1] B] (π→[n + 1] g) q :=
@ -202,14 +191,7 @@ namespace eq
apply ap tr, apply apn_con apply ap tr, apply apn_con
end end
definition homotopy_group_homomorphism [constructor] (n : ) [H : is_succ n] {A B : Type*} /- todo: rename πg→ -/
(f : A →* B) : πg[n] A →g πg[n] B :=
begin
induction H with n, fconstructor,
{ exact homotopy_group_functor (n+1) f},
{ apply homotopy_group_functor_mul}
end
notation `π→g[`:95 n:0 `]`:0 := homotopy_group_homomorphism n notation `π→g[`:95 n:0 `]`:0 := homotopy_group_homomorphism n
definition homotopy_group_homomorphism_pcompose (n : ) [H : is_succ n] {A B C : Type*} (g : B →* C) definition homotopy_group_homomorphism_pcompose (n : ) [H : is_succ n] {A B C : Type*} (g : B →* C)
@ -218,12 +200,10 @@ namespace eq
induction H with n, exact to_homotopy (homotopy_group_functor_pcompose (succ n) g f) induction H with n, exact to_homotopy (homotopy_group_functor_pcompose (succ n) g f)
end end
/- todo: use is_succ -/
definition homotopy_group_isomorphism_of_pequiv [constructor] (n : ) {A B : Type*} (f : A ≃* B) definition homotopy_group_isomorphism_of_pequiv [constructor] (n : ) {A B : Type*} (f : A ≃* B)
: πg[n+1] A ≃g πg[n+1] B := : πg[n+1] A ≃g πg[n+1] B :=
begin gtrunc_isomorphism_gtrunc (gloopn_isomorphism (n+1) f)
apply isomorphism.mk (homotopy_group_homomorphism (succ n) f),
exact is_equiv_homotopy_group_functor _ _ _,
end
definition homotopy_group_add (A : Type*) (n m : ) : definition homotopy_group_add (A : Type*) (n m : ) :
πg[n+m+1] A ≃g πg[n+1] (Ω[m] A) := πg[n+m+1] A ≃g πg[n+1] (Ω[m] A) :=

View file

@ -1,8 +1,7 @@
/- /-
Copyright (c) 2014 Jeremy Avigad. All rights reserved. Copyright (c) 2014 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn
-/ -/
import algebra.binary algebra.priority import algebra.binary algebra.priority
@ -327,6 +326,9 @@ end inf_group
structure ab_inf_group [class] (A : Type) extends inf_group A, comm_inf_monoid A structure ab_inf_group [class] (A : Type) extends inf_group A, comm_inf_monoid A
theorem mul.comm4 [s : ab_inf_group A] (a b c d : A) : (a * b) * (c * d) = (a * c) * (b * d) :=
binary.comm4 mul.comm mul.assoc a b c d
/- additive inf_group -/ /- additive inf_group -/
definition add_inf_group [class] : Type → Type := inf_group definition add_inf_group [class] : Type → Type := inf_group

View file

@ -0,0 +1,386 @@
/-
Copyright (c) 2018 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
-/
import .bundled .homomorphism types.nat.hott
open algebra eq is_equiv equiv pointed function is_trunc nat
universe variable u
namespace group
/- left and right actions -/
definition is_equiv_mul_right_inf [constructor] {A : InfGroup} (a : A) : is_equiv (λb, b * a) :=
adjointify _ (λb : A, b * a⁻¹) (λb, !inv_mul_cancel_right) (λb, !mul_inv_cancel_right)
definition right_action_inf [constructor] {A : InfGroup} (a : A) : A ≃ A :=
equiv.mk _ (is_equiv_mul_right_inf a)
/- homomorphisms -/
structure inf_homomorphism (G₁ G₂ : InfGroup) : Type :=
(φ : G₁ → G₂)
(p : is_mul_hom φ)
infix ` →∞g `:55 := inf_homomorphism
abbreviation inf_group_fun [unfold 3] [coercion] [reducible] := @inf_homomorphism.φ
definition inf_homomorphism.struct [unfold 3] [instance] [priority 900] {G₁ G₂ : InfGroup}
(φ : G₁ →∞g G₂) : is_mul_hom φ :=
inf_homomorphism.p φ
variables {G G₁ G₂ G₃ : InfGroup} {g h : G₁} {ψ : G₂ →∞g G₃} {φ₁ φ₂ : G₁ →∞g G₂} (φ : G₁ →∞g G₂)
definition to_respect_mul_inf /- φ -/ (g h : G₁) : φ (g * h) = φ g * φ h :=
respect_mul φ g h
theorem to_respect_one_inf /- φ -/ : φ 1 = 1 :=
have φ 1 * φ 1 = φ 1 * 1, by rewrite [-to_respect_mul_inf φ, +mul_one],
eq_of_mul_eq_mul_left' this
theorem to_respect_inv_inf /- φ -/ (g : G₁) : φ g⁻¹ = (φ g)⁻¹ :=
have φ (g⁻¹) * φ g = 1, by rewrite [-to_respect_mul_inf φ, mul.left_inv, to_respect_one_inf φ],
eq_inv_of_mul_eq_one this
definition pmap_of_inf_homomorphism [constructor] /- φ -/ : G₁ →* G₂ :=
pmap.mk φ begin exact to_respect_one_inf φ end
definition inf_homomorphism_change_fun [constructor] {G₁ G₂ : InfGroup}
(φ : G₁ →∞g G₂) (f : G₁ → G₂) (p : φ ~ f) : G₁ →∞g G₂ :=
inf_homomorphism.mk f
(λg h, (p (g * h))⁻¹ ⬝ to_respect_mul_inf φ g h ⬝ ap011 mul (p g) (p h))
/- categorical structure of groups + homomorphisms -/
definition inf_homomorphism_compose [constructor] [trans] [reducible]
(ψ : G₂ →∞g G₃) (φ : G₁ →∞g G₂) : G₁ →∞g G₃ :=
inf_homomorphism.mk (ψ ∘ φ) (is_mul_hom_compose _ _)
variable (G)
definition inf_homomorphism_id [constructor] [refl] : G →∞g G :=
inf_homomorphism.mk (@id G) (is_mul_hom_id G)
variable {G}
abbreviation inf_gid [constructor] := @inf_homomorphism_id
infixr ` ∘∞g `:75 := inf_homomorphism_compose
structure inf_isomorphism (A B : InfGroup) :=
(to_hom : A →∞g B)
(is_equiv_to_hom : is_equiv to_hom)
infix ` ≃∞g `:25 := inf_isomorphism
attribute inf_isomorphism.to_hom [coercion]
attribute inf_isomorphism.is_equiv_to_hom [instance]
attribute inf_isomorphism._trans_of_to_hom [unfold 3]
definition equiv_of_inf_isomorphism [constructor] (φ : G₁ ≃∞g G₂) : G₁ ≃ G₂ :=
equiv.mk φ _
definition pequiv_of_inf_isomorphism [constructor] (φ : G₁ ≃∞g G₂) :
G₁ ≃* G₂ :=
pequiv.mk φ begin esimp, exact _ end begin esimp, exact to_respect_one_inf φ end
definition inf_isomorphism_of_equiv [constructor] (φ : G₁ ≃ G₂)
(p : Πg₁ g₂, φ (g₁ * g₂) = φ g₁ * φ g₂) : G₁ ≃∞g G₂ :=
inf_isomorphism.mk (inf_homomorphism.mk φ p) !to_is_equiv
definition inf_isomorphism_of_eq [constructor] {G₁ G₂ : InfGroup} (φ : G₁ = G₂) : G₁ ≃∞g G₂ :=
inf_isomorphism_of_equiv (equiv_of_eq (ap InfGroup.carrier φ))
begin intros, induction φ, reflexivity end
definition to_ginv_inf [constructor] (φ : G₁ ≃∞g G₂) : G₂ →∞g G₁ :=
inf_homomorphism.mk φ⁻¹
abstract begin
intro g₁ g₂, apply inj' φ,
rewrite [respect_mul φ, +right_inv φ]
end end
variable (G)
definition inf_isomorphism.refl [refl] [constructor] : G ≃∞g G :=
inf_isomorphism.mk !inf_gid !is_equiv_id
variable {G}
definition inf_isomorphism.symm [symm] [constructor] (φ : G₁ ≃∞g G₂) : G₂ ≃∞g G₁ :=
inf_isomorphism.mk (to_ginv_inf φ) !is_equiv_inv
definition inf_isomorphism.trans [trans] [constructor] (φ : G₁ ≃∞g G₂) (ψ : G₂ ≃∞g G₃) :
G₁ ≃∞g G₃ :=
inf_isomorphism.mk (ψ ∘∞g φ) (is_equiv_compose ψ φ _ _)
definition inf_isomorphism.eq_trans [trans] [constructor]
{G₁ G₂ : InfGroup} {G₃ : InfGroup} (φ : G₁ = G₂) (ψ : G₂ ≃∞g G₃) : G₁ ≃∞g G₃ :=
proof inf_isomorphism.trans (inf_isomorphism_of_eq φ) ψ qed
definition inf_isomorphism.trans_eq [trans] [constructor]
{G₁ : InfGroup} {G₂ G₃ : InfGroup} (φ : G₁ ≃∞g G₂) (ψ : G₂ = G₃) : G₁ ≃∞g G₃ :=
inf_isomorphism.trans φ (inf_isomorphism_of_eq ψ)
postfix `⁻¹ᵍ⁸`:(max + 1) := inf_isomorphism.symm
infixl ` ⬝∞g `:75 := inf_isomorphism.trans
infixl ` ⬝∞gp `:75 := inf_isomorphism.trans_eq
infixl ` ⬝∞pg `:75 := inf_isomorphism.eq_trans
definition pmap_of_inf_isomorphism [constructor] (φ : G₁ ≃∞g G₂) : G₁ →* G₂ :=
pequiv_of_inf_isomorphism φ
definition to_fun_inf_isomorphism_trans {G H K : InfGroup} (φ : G ≃∞g H) (ψ : H ≃∞g K) :
φ ⬝∞g ψ ~ ψ ∘ φ :=
by reflexivity
definition inf_homomorphism_mul [constructor] {G H : AbInfGroup} (φ ψ : G →∞g H) : G →∞g H :=
inf_homomorphism.mk (λg, φ g * ψ g)
abstract begin
intro g g', refine ap011 mul !to_respect_mul_inf !to_respect_mul_inf ⬝ _,
refine !mul.assoc ⬝ ap (mul _) (!mul.assoc⁻¹ ⬝ ap (λx, x * _) !mul.comm ⬝ !mul.assoc) ⬝
!mul.assoc⁻¹
end end
definition trivial_inf_homomorphism (A B : InfGroup) : A →∞g B :=
inf_homomorphism.mk (λa, 1) (λa a', (mul_one 1)⁻¹)
/- given an equivalence A ≃ B we can transport a group structure on A to a group structure on B -/
section
parameters {A B : Type} (f : A ≃ B) (H : inf_group A)
include H
definition inf_group_equiv_mul (b b' : B) : B := f (f⁻¹ᶠ b * f⁻¹ᶠ b')
definition inf_group_equiv_one : B := f one
definition inf_group_equiv_inv (b : B) : B := f (f⁻¹ᶠ b)⁻¹
local infix * := inf_group_equiv_mul
local postfix ^ := inf_group_equiv_inv
local notation 1 := inf_group_equiv_one
theorem inf_group_equiv_mul_assoc (b₁ b₂ b₃ : B) : (b₁ * b₂) * b₃ = b₁ * (b₂ * b₃) :=
by rewrite [↑inf_group_equiv_mul, +left_inv f, mul.assoc]
theorem inf_group_equiv_one_mul (b : B) : 1 * b = b :=
by rewrite [↑inf_group_equiv_mul, ↑inf_group_equiv_one, left_inv f, one_mul, right_inv f]
theorem inf_group_equiv_mul_one (b : B) : b * 1 = b :=
by rewrite [↑inf_group_equiv_mul, ↑inf_group_equiv_one, left_inv f, mul_one, right_inv f]
theorem inf_group_equiv_mul_left_inv (b : B) : b^ * b = 1 :=
by rewrite [↑inf_group_equiv_mul, ↑inf_group_equiv_one, ↑inf_group_equiv_inv,
+left_inv f, mul.left_inv]
definition inf_group_equiv_closed : inf_group B :=
⦃inf_group,
mul := inf_group_equiv_mul,
mul_assoc := inf_group_equiv_mul_assoc,
one := inf_group_equiv_one,
one_mul := inf_group_equiv_one_mul,
mul_one := inf_group_equiv_mul_one,
inv := inf_group_equiv_inv,
mul_left_inv := inf_group_equiv_mul_left_inv⦄
end
section
variables {A B : Type} (f : A ≃ B) (H : ab_inf_group A)
include H
definition inf_group_equiv_mul_comm (b b' : B) :
inf_group_equiv_mul f _ b b' = inf_group_equiv_mul f _ b' b :=
by rewrite [↑inf_group_equiv_mul, mul.comm]
definition ab_inf_group_equiv_closed : ab_inf_group B :=
⦃ ab_inf_group, inf_group_equiv_closed f _, mul_comm := inf_group_equiv_mul_comm f H ⦄
end
variable (G)
/- the trivial ∞-group -/
open unit
definition inf_group_unit [constructor] : inf_group unit :=
inf_group.mk (λx y, star) (λx y z, idp) star (unit.rec idp) (unit.rec idp) (λx, star) (λx, idp)
definition ab_inf_group_unit [constructor] : ab_inf_group unit :=
⦃ab_inf_group, inf_group_unit, mul_comm := λx y, idp⦄
definition trivial_inf_group [constructor] : InfGroup :=
InfGroup.mk _ inf_group_unit
definition AbInfGroup_of_InfGroup (G : InfGroup.{u}) (H : Π x y : G, x * y = y * x) :
AbInfGroup.{u} :=
begin
induction G,
fapply AbInfGroup.mk,
assumption,
exact ⦃ab_inf_group, struct', mul_comm := H⦄
end
definition trivial_ab_inf_group : AbInfGroup.{0} :=
begin
fapply AbInfGroup_of_InfGroup trivial_inf_group, intro x y, reflexivity
end
definition trivial_inf_group_of_is_contr [H : is_contr G] : G ≃∞g trivial_inf_group :=
begin
fapply inf_isomorphism_of_equiv,
{ exact equiv_unit_of_is_contr _ _ },
{ intros, reflexivity}
end
definition ab_inf_group_of_is_contr (A : Type) (H : is_contr A) : ab_inf_group A :=
have ab_inf_group unit, from ab_inf_group_unit,
ab_inf_group_equiv_closed (equiv_unit_of_is_contr A _)⁻¹ᵉ _
definition inf_group_of_is_contr (A : Type) (H : is_contr A) : inf_group A :=
have ab_inf_group A, from ab_inf_group_of_is_contr A H, by exact _
definition ab_inf_group_lift_unit : ab_inf_group (lift unit) :=
ab_inf_group_of_is_contr (lift unit) _
definition trivial_ab_inf_group_lift : AbInfGroup :=
AbInfGroup.mk _ ab_inf_group_lift_unit
definition from_trivial_ab_inf_group (A : AbInfGroup) : trivial_ab_inf_group →∞g A :=
trivial_inf_homomorphism trivial_ab_inf_group A
definition to_trivial_ab_inf_group (A : AbInfGroup) : A →∞g trivial_ab_inf_group :=
trivial_inf_homomorphism A trivial_ab_inf_group
/- infinity pgroups are infgroups where 1 is definitionally the point of X -/
structure inf_pgroup [class] (X : Type*) extends inf_semigroup X, has_inv X :=
(pt_mul : Πa, mul pt a = a)
(mul_pt : Πa, mul a pt = a)
(mul_left_inv_pt : Πa, mul (inv a) a = pt)
definition pt_mul (X : Type*) [inf_pgroup X] (x : X) : pt * x = x := inf_pgroup.pt_mul x
definition mul_pt (X : Type*) [inf_pgroup X] (x : X) : x * pt = x := inf_pgroup.mul_pt x
definition mul_left_inv_pt (X : Type*) [inf_pgroup X] (x : X) : x⁻¹ * x = pt :=
inf_pgroup.mul_left_inv_pt x
definition inf_group_of_inf_pgroup [reducible] [instance] (X : Type*) [H : inf_pgroup X]
: inf_group X :=
⦃inf_group, H,
one := pt,
one_mul := inf_pgroup.pt_mul ,
mul_one := inf_pgroup.mul_pt,
mul_left_inv := inf_pgroup.mul_left_inv_pt⦄
definition inf_pgroup_of_inf_group (X : Type*) [H : inf_group X] (p : one = pt :> X) :
inf_pgroup X :=
begin
cases X with X x, esimp at *, induction p,
exact ⦃inf_pgroup, H,
pt_mul := one_mul,
mul_pt := mul_one,
mul_left_inv_pt := mul.left_inv⦄
end
definition inf_Group_of_inf_pgroup (G : Type*) [inf_pgroup G] : InfGroup :=
InfGroup.mk G _
definition inf_pgroup_InfGroup [instance] (G : InfGroup) : inf_pgroup G :=
⦃ inf_pgroup, InfGroup.struct G,
pt_mul := one_mul,
mul_pt := mul_one,
mul_left_inv_pt := mul.left_inv ⦄
section
parameters {A B : Type*} (f : A ≃* B) (s : inf_pgroup A)
include s
definition inf_pgroup_pequiv_mul (b b' : B) : B := f (f⁻¹ᶠ b * f⁻¹ᶠ b')
definition inf_pgroup_pequiv_inv (b : B) : B := f (f⁻¹ᶠ b)⁻¹
local infix * := inf_pgroup_pequiv_mul
local postfix ^ := inf_pgroup_pequiv_inv
theorem inf_pgroup_pequiv_mul_assoc (b₁ b₂ b₃ : B) : (b₁ * b₂) * b₃ = b₁ * (b₂ * b₃) :=
begin
refine ap (λa, f (mul a _)) (left_inv f _) ⬝ _ ⬝ ap (λa, f (mul _ a)) (left_inv f _)⁻¹,
exact ap f !mul.assoc
end
theorem inf_pgroup_pequiv_pt_mul (b : B) : pt * b = b :=
by rewrite [↑inf_pgroup_pequiv_mul, respect_pt, pt_mul]; apply right_inv f
theorem inf_pgroup_pequiv_mul_pt (b : B) : b * pt = b :=
by rewrite [↑inf_pgroup_pequiv_mul, respect_pt, mul_pt]; apply right_inv f
theorem inf_pgroup_pequiv_mul_left_inv_pt (b : B) : b^ * b = pt :=
begin
refine ap (λa, f (mul a _)) (left_inv f _) ⬝ _,
exact ap f !mul_left_inv_pt ⬝ !respect_pt,
end
definition inf_pgroup_pequiv_closed : inf_pgroup B :=
⦃inf_pgroup,
mul := inf_pgroup_pequiv_mul,
mul_assoc := inf_pgroup_pequiv_mul_assoc,
pt_mul := inf_pgroup_pequiv_pt_mul,
mul_pt := inf_pgroup_pequiv_mul_pt,
inv := inf_pgroup_pequiv_inv,
mul_left_inv_pt := inf_pgroup_pequiv_mul_left_inv_pt⦄
end
/- infgroup from loop spaces -/
definition inf_pgroup_loop [constructor] [instance] (A : Type*) : inf_pgroup (Ω A) :=
inf_pgroup.mk concat con.assoc inverse idp_con con_idp con.left_inv
definition inf_group_loop [constructor] (A : Type*) : inf_group (Ω A) := _
definition ab_inf_group_loop [constructor] [instance] (A : Type*) : ab_inf_group (Ω (Ω A)) :=
⦃ab_inf_group, inf_group_loop _, mul_comm := eckmann_hilton⦄
definition inf_group_loopn (n : ) (A : Type*) [H : is_succ n] : inf_group (Ω[n] A) :=
by induction H; exact _
definition ab_inf_group_loopn (n : ) (A : Type*) [H : is_at_least_two n] :
ab_inf_group (Ω[n] A) :=
by induction H; exact _
definition gloop [constructor] (A : Type*) : InfGroup :=
InfGroup.mk (Ω A) (inf_group_loop A)
definition gloopn (n : ) [H : is_succ n] (A : Type*) : InfGroup :=
InfGroup.mk (Ω[n] A) (inf_group_loopn n A)
definition agloopn (n : ) [H : is_at_least_two n] (A : Type*) : AbInfGroup :=
AbInfGroup.mk (Ω[n] A) (ab_inf_group_loopn n A)
definition gloopn' (n : ) (A : InfGroup) : InfGroup :=
InfGroup.mk (Ω[n] A) (by cases n; exact InfGroup.struct A; apply inf_group_loopn)
notation `Ωg` := gloop
notation `Ωg[`:95 n:0 `]`:0 := gloopn n
notation `Ωag[`:95 n:0 `]`:0 := agloopn n
notation `Ωg'[`:95 n:0 `]`:0 := gloopn' n
definition gap1 {A B : Type*} (f : A →* B) : Ωg A →∞g Ωg B :=
inf_homomorphism.mk (Ω→ f) (ap1_con f)
definition gapn (n : ) [H : is_succ n] {A B : Type*} (f : A →* B) : Ωg[n] A →∞g Ωg[n] B :=
inf_homomorphism.mk (Ω→[n] f) (by induction H with n; exact apn_con n f)
notation `Ωg→` := gap1
notation `Ωg→[`:95 n:0 `]`:0 := gapn n
definition gloop_isomorphism {A B : Type*} (f : A ≃* B) : Ωg A ≃∞g Ωg B :=
inf_isomorphism.mk (Ωg→ f) (to_is_equiv (loop_pequiv_loop f))
definition gloopn_isomorphism (n : ) [H : is_succ n] {A B : Type*} (f : A ≃* B) :
Ωg[n] A ≃∞g Ωg[n] B :=
inf_isomorphism.mk (Ωg→[n] f) (to_is_equiv (loopn_pequiv_loopn n f))
notation `Ωg≃` := gloop_isomorphism
notation `Ωg≃[`:95 n:0 `]`:0 := gloopn_isomorphism
definition gloopn_succ_in (n : ) [H : is_succ n] (A : Type*) : Ωg[succ n] A ≃∞g Ωg[n] (Ω A) :=
inf_isomorphism_of_equiv (loopn_succ_in n A) (by induction H with n; exact loopn_succ_in_con)
end group

View file

@ -1,7 +1,7 @@
/- /-
Copyright (c) 2014 Jeremy Avigad. All rights reserved. Copyright (c) 2014 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn
Structures with multiplicative and additive components, including semirings, rings, and fields. Structures with multiplicative and additive components, including semirings, rings, and fields.
The development is modeled after Isabelle's library. The development is modeled after Isabelle's library.

View file

@ -64,7 +64,7 @@ namespace algebra
end end
parameter (A) parameter (A)
definition trunc_inf_group [constructor] [instance] : inf_group (trunc n A) := definition inf_group_trunc [constructor] [instance] : inf_group (trunc n A) :=
⦃inf_group, ⦃inf_group,
mul := algebra.trunc_mul n, mul := algebra.trunc_mul n,
mul_assoc := algebra.trunc_mul_assoc n, mul_assoc := algebra.trunc_mul_assoc n,
@ -74,11 +74,17 @@ namespace algebra
inv := algebra.trunc_inv n, inv := algebra.trunc_inv n,
mul_left_inv := algebra.trunc_mul_left_inv n⦄ mul_left_inv := algebra.trunc_mul_left_inv n⦄
definition trunc_group [constructor] : group (trunc 0 A) := definition group_trunc [constructor] : group (trunc 0 A) :=
group_of_inf_group _ group_of_inf_group _
end end
definition igtrunc [constructor] (n : ℕ₋₂) (A : InfGroup) : InfGroup :=
InfGroup.mk (trunc n A) (inf_group_trunc n A)
definition gtrunc [constructor] (A : InfGroup) : Group :=
Group.mk (trunc 0 A) (group_trunc A)
section section
variables (n : trunc_index) {A : Type} [ab_inf_group A] variables (n : trunc_index) {A : Type} [ab_inf_group A]
@ -90,12 +96,18 @@ namespace algebra
end end
variable (A) variable (A)
definition trunc_ab_inf_group [constructor] [instance] : ab_inf_group (trunc n A) := definition ab_inf_group_trunc [constructor] [instance] : ab_inf_group (trunc n A) :=
⦃ab_inf_group, trunc_inf_group n A, mul_comm := algebra.trunc_mul_comm n⦄ ⦃ab_inf_group, inf_group_trunc n A, mul_comm := algebra.trunc_mul_comm n⦄
definition trunc_ab_group [constructor] : ab_group (trunc 0 A) := definition ab_group_trunc [constructor] : ab_group (trunc 0 A) :=
ab_group_of_ab_inf_group _ ab_group_of_ab_inf_group _
definition aigtrunc [constructor] (n : ℕ₋₂) (A : AbInfGroup) : AbInfGroup :=
AbInfGroup.mk (trunc n A) (ab_inf_group_trunc n A)
definition agtrunc [constructor] (A : AbInfGroup) : AbGroup :=
AbGroup.mk (trunc 0 A) (ab_group_trunc A)
end end
end algebra end algebra

View file

@ -25,6 +25,7 @@ namespace choice
equiv_of_is_prop equiv_of_is_prop
(λ H X A P HX HA HP HI, trunc_functor _ (to_fun !sigma_pi_equiv_pi_sigma) (H X A P HX HA HP HI)) (λ H X A P HX HA HP HI, trunc_functor _ (to_fun !sigma_pi_equiv_pi_sigma) (H X A P HX HA HP HI))
(λ H X A P HX HA HP HI, trunc_functor _ (to_inv !sigma_pi_equiv_pi_sigma) (H X A P HX HA HP HI)) (λ H X A P HX HA HP HI, trunc_functor _ (to_inv !sigma_pi_equiv_pi_sigma) (H X A P HX HA HP HI))
_ _
-- AC_cart can be derived from AC' by setting P := λ _ _ , unit. -- AC_cart can be derived from AC' by setting P := λ _ _ , unit.
definition AC_cart_of_AC' : AC'.{u} -> AC_cart.{u} := definition AC_cart_of_AC' : AC'.{u} -> AC_cart.{u} :=
@ -39,7 +40,7 @@ namespace choice
-- Which is enough to show AC' ≃ AC_cart, since both are props. -- Which is enough to show AC' ≃ AC_cart, since both are props.
definition AC'_equiv_AC_cart : AC'.{u} ≃ AC_cart.{u} := definition AC'_equiv_AC_cart : AC'.{u} ≃ AC_cart.{u} :=
equiv_of_is_prop AC_cart_of_AC'.{u} AC'_of_AC_cart.{u} equiv_of_is_prop AC_cart_of_AC'.{u} AC'_of_AC_cart.{u} _ _
-- 3.8.2. AC ≃ AC_cart follows by transitivity. -- 3.8.2. AC ≃ AC_cart follows by transitivity.
definition AC_equiv_AC_cart : AC.{u} ≃ AC_cart.{u} := definition AC_equiv_AC_cart : AC.{u} ≃ AC_cart.{u} :=

View file

@ -20,6 +20,8 @@ definition image [constructor] (f : A → B) (b : B) : Prop := Prop.mk (image' f
definition total_image {A B : Type} (f : A → B) : Type := sigma (image f) definition total_image {A B : Type} (f : A → B) : Type := sigma (image f)
/- properties of functions -/
definition is_embedding [class] (f : A → B) := Π(a a' : A), is_equiv (ap f : a = a' → f a = f a') definition is_embedding [class] (f : A → B) := Π(a a' : A), is_equiv (ap f : a = a' → f a = f a')
definition is_surjective [class] (f : A → B) := Π(b : B), image f b definition is_surjective [class] (f : A → B) := Π(b : B), image f b
@ -308,9 +310,9 @@ namespace function
definition is_embedding_compose (g : B → C) (f : A → B) definition is_embedding_compose (g : B → C) (f : A → B)
(H₁ : is_embedding g) (H₂ : is_embedding f) : is_embedding (g ∘ f) := (H₁ : is_embedding g) (H₂ : is_embedding f) : is_embedding (g ∘ f) :=
begin begin
intros, apply @(is_equiv.homotopy_closed (ap g ∘ ap f)), intros, apply is_equiv.homotopy_closed (ap g ∘ ap f),
{ apply is_equiv_compose}, { symmetry, exact ap_compose g f },
symmetry, exact ap_compose g f { exact is_equiv_compose _ _ _ _ }
end end
definition is_surjective_compose (g : B → C) (f : A → B) definition is_surjective_compose (g : B → C) (f : A → B)
@ -361,7 +363,7 @@ namespace function
{a a' : A} {b b' : B} (q : f a = b) (q' : f a' = b') : is_equiv (ap1_gen f q q') := {a a' : A} {b b' : B} (q : f a = b) (q' : f a' = b') : is_equiv (ap1_gen f q q') :=
begin begin
induction q, induction q', induction q, induction q',
exact is_equiv.homotopy_closed _ (ap1_gen_idp_left f)⁻¹ʰᵗʸ, exact is_equiv.homotopy_closed _ (ap1_gen_idp_left f)⁻¹ʰᵗʸ _,
end end
definition is_equiv_ap1_of_is_embedding {A B : Type*} (f : A →* B) [is_embedding f] : definition is_equiv_ap1_of_is_embedding {A B : Type*} (f : A →* B) [is_embedding f] :

View file

@ -144,7 +144,7 @@ namespace one_step_tr
theorem trunc_0_one_step_tr_equiv (A : Type) : trunc 0 (one_step_tr A) ≃ ∥ A ∥ := theorem trunc_0_one_step_tr_equiv (A : Type) : trunc 0 (one_step_tr A) ≃ ∥ A ∥ :=
begin begin
apply equiv_of_is_prop, refine equiv_of_is_prop _ _ _ _,
{ intro x, refine trunc.rec _ x, clear x, intro x, induction x, { intro x, refine trunc.rec _ x, clear x, intro x, induction x,
{ exact trunc.tr a}, { exact trunc.tr a},
{ apply is_prop.elim}}, { apply is_prop.elim}},

View file

@ -98,9 +98,9 @@ namespace set_quotient
definition is_trunc_set_quotient [instance] (n : ℕ₋₂) {A : Type} (R : A → A → Prop) [is_trunc n A] : definition is_trunc_set_quotient [instance] (n : ℕ₋₂) {A : Type} (R : A → A → Prop) [is_trunc n A] :
is_trunc n (set_quotient R) := is_trunc n (set_quotient R) :=
begin begin
cases n with n, { apply is_contr_of_inhabited_prop, exact class_of !center }, cases n with n, { refine is_contr_of_inhabited_prop _ _, exact class_of !center },
cases n with n, { apply _ }, cases n with n, { apply _ },
apply is_trunc_succ_succ_of_is_set exact is_trunc_succ_succ_of_is_set _ _ _
end end
definition is_equiv_class_of [constructor] {A : Type} [is_set A] (R : A → A → Prop) definition is_equiv_class_of [constructor] {A : Type} [is_set A] (R : A → A → Prop)
@ -135,10 +135,8 @@ namespace set_quotient
: Prop := : Prop :=
set_quotient.elim_on x (R a) set_quotient.elim_on x (R a)
begin begin
intros a' a'' H1, intros a' a'' H1, apply tua,
refine to_inv !trunctype_eq_equiv _, esimp, refine equiv_of_is_prop _ _ _ _,
apply ua,
apply equiv_of_is_prop,
{ intro H2, exact is_transitive.trans R H2 H1}, { intro H2, exact is_transitive.trans R H2 H1},
{ intro H2, apply is_transitive.trans R H2, exact is_symmetric.symm R H1} { intro H2, apply is_transitive.trans R H2, exact is_symmetric.symm R H1}
end end

View file

@ -36,11 +36,11 @@ namespace trunc
local attribute is_trunc_eq [instance] local attribute is_trunc_eq [instance]
variables {A n} variables {n A}
definition untrunc_of_is_trunc [reducible] [unfold 4] [H : is_trunc n A] : trunc n A → A := definition untrunc_of_is_trunc [reducible] [unfold 4] [H : is_trunc n A] : trunc n A → A :=
trunc.rec id trunc.rec id
variables (A n) variables (n A)
definition is_equiv_tr [instance] [constructor] [H : is_trunc n A] : is_equiv (@tr n A) := definition is_equiv_tr [instance] [constructor] [H : is_trunc n A] : is_equiv (@tr n A) :=
adjointify _ adjointify _
(untrunc_of_is_trunc) (untrunc_of_is_trunc)
@ -51,7 +51,7 @@ namespace trunc
(equiv.mk tr _)⁻¹ᵉ (equiv.mk tr _)⁻¹ᵉ
definition is_trunc_of_is_equiv_tr [H : is_equiv (@tr n A)] : is_trunc n A := definition is_trunc_of_is_equiv_tr [H : is_equiv (@tr n A)] : is_trunc n A :=
is_trunc_is_equiv_closed n (@tr n _)⁻¹ is_trunc_is_equiv_closed n (@tr n _)⁻¹ᶠ _ _
/- Functoriality -/ /- Functoriality -/
definition trunc_functor [unfold 5] (f : X → Y) : trunc n X → trunc n Y := definition trunc_functor [unfold 5] (f : X → Y) : trunc n X → trunc n Y :=
@ -131,7 +131,7 @@ namespace trunc
trunc.elim (sigma.rec H') H trunc.elim (sigma.rec H') H
definition is_contr_of_merely_prop [H : is_prop A] (aa : merely A) : is_contr A := definition is_contr_of_merely_prop [H : is_prop A] (aa : merely A) : is_contr A :=
is_contr_of_inhabited_prop (trunc.rec_on aa id) is_contr_of_inhabited_prop (trunc.rec_on aa id) _
section section
open sigma.ops open sigma.ops

View file

@ -129,59 +129,56 @@ namespace EM
is_conn_EM1' G is_conn_EM1' G
variable {G} variable {G}
definition EM1_map [unfold 6] {G : Group} {X : Type*} (e : G → Ω X) open infgroup
(r : Πg h, e (g * h) = e g ⬝ e h) [is_trunc 1 X] : EM1 G → X := definition EM1_map [unfold 6] {G : Group} {X : Type*} (e : G →∞g Ωg X) [is_trunc 1 X] :
EM1 G → X :=
begin begin
intro x, induction x using EM.elim, intro x, induction x using EM.elim,
{ exact Point X }, { exact Point X },
{ exact e g }, { exact e g },
{ exact r g h } { exact to_respect_mul_inf e g h }
end end
/- Uniqueness of K(G, 1) -/ /- Uniqueness of K(G, 1) -/
definition EM1_pmap [constructor] {G : Group} {X : Type*} (e : G → Ω X) definition EM1_pmap [constructor] {G : Group} {X : Type*} (e : G →∞g Ωg X) [is_trunc 1 X] :
(r : Πg h, e (g * h) = e g ⬝ e h) [is_trunc 1 X] : EM1 G →* X := EM1 G →* X :=
pmap.mk (EM1_map e r) idp pmap.mk (EM1_map e) idp
variable (G) variable (G)
definition loop_EM1 [constructor] : G ≃* Ω (EM1 G) := definition loop_EM1 [constructor] : G ≃* Ω (EM1 G) :=
(pequiv_of_equiv (base_eq_base_equiv G) idp)⁻¹ᵉ* (pequiv_of_equiv (base_eq_base_equiv G) idp)⁻¹ᵉ*
variable {G} variable {G}
definition loop_EM1_pmap {G : Group} {X : Type*} (e : G →* Ω X) definition loop_EM1_pmap {G : Group} {X : Type*} (e : G →∞g Ωg X) [is_trunc 1 X] :
(r : Πg h, e (g * h) = e g ⬝ e h) [is_trunc 1 X] : Ω→(EM1_pmap e r) ∘* loop_EM1 G ~* e := Ω→(EM1_pmap e) ∘* loop_EM1 G ~* pmap_of_inf_homomorphism e :=
begin begin
fapply phomotopy.mk, fapply phomotopy.mk,
{ intro g, refine !idp_con ⬝ elim_pth r g }, { intro g, refine !idp_con ⬝ elim_pth (to_respect_mul_inf e) g },
{ apply is_set.elim } { apply is_set.elim }
end end
definition EM1_pequiv'.{u} {G : Group.{u}} {X : pType.{u}} (e : G ≃* Ω X) definition EM1_pequiv'.{u} {G : Group.{u}} {X : pType.{u}} (e : G ≃∞g Ωg X)
(r : Πg h, e (g * h) = e g ⬝ e h) [is_conn 0 X] [is_trunc 1 X] : EM1 G ≃* X := [is_conn 0 X] [is_trunc 1 X] : EM1 G ≃* X :=
begin begin
apply pequiv_of_pmap (EM1_pmap e r), apply pequiv_of_pmap (EM1_pmap e),
apply whitehead_principle_pointed 1, apply whitehead_principle_pointed 1,
intro k, cases k with k, intro k, cases k with k,
{ apply @is_equiv_of_is_contr, { apply @is_equiv_of_is_contr,
all_goals (esimp; exact _)}, all_goals (esimp; exact _)},
{ cases k with k, { cases k with k,
{ apply is_equiv_trunc_functor, esimp, { apply is_equiv_trunc_functor, esimp,
apply is_equiv.homotopy_closed, rotate 1, apply is_equiv.homotopy_closed,
{ symmetry, exact phomotopy_pinv_right_of_phomotopy (loop_EM1_pmap _ _) }, { symmetry, exact phomotopy_pinv_right_of_phomotopy (loop_EM1_pmap e) },
apply is_equiv_compose e, apply pequiv.to_is_equiv }, refine is_equiv_compose e _ _ _, apply inf_isomorphism.is_equiv_to_hom },
{ apply @is_equiv_of_is_contr, { apply @is_equiv_of_is_contr,
do 2 exact trivial_homotopy_group_of_is_trunc _ (succ_lt_succ !zero_lt_succ)}} do 2 exact trivial_homotopy_group_of_is_trunc _ (succ_lt_succ !zero_lt_succ)}}
end end
definition EM1_pequiv.{u} {G : Group.{u}} {X : pType.{u}} (e : G ≃g π₁ X) definition EM1_pequiv.{u} {G : Group.{u}} {X : pType.{u}} (e : G ≃g π₁ X)
[is_conn 0 X] [is_trunc 1 X] : EM1 G ≃* X := [is_conn 0 X] [is_trunc 1 X] : EM1 G ≃* X :=
begin have is_set (Ωg X), from !is_trunc_loop,
apply EM1_pequiv' (pequiv_of_isomorphism e ⬝e* ptrunc_pequiv 0 (Ω X)), EM1_pequiv' (inf_isomorphism_of_isomorphism e ⬝∞g gtrunc_isomorphism (Ωg X))
refine equiv.preserve_binary_of_inv_preserve _ mul concat _,
intro p q,
exact to_respect_mul e⁻¹ᵍ (tr p) (tr q)
end
definition EM1_pequiv_type (X : Type*) [is_conn 0 X] [is_trunc 1 X] : EM1 (π₁ X) ≃* X := definition EM1_pequiv_type (X : Type*) [is_conn 0 X] [is_trunc 1 X] : EM1 (π₁ X) ≃* X :=
EM1_pequiv !isomorphism.refl EM1_pequiv !isomorphism.refl
@ -282,41 +279,30 @@ namespace EM
!loopn_succ_in⁻¹ᵉ* ∘* apn (succ n) !loop_EMadd1 ∘* loopn_EMadd1 G n := !loopn_succ_in⁻¹ᵉ* ∘* apn (succ n) !loop_EMadd1 ∘* loopn_EMadd1 G n :=
by reflexivity by reflexivity
definition EM_up {G : AbGroup} {X : Type*} {n : } (e : G →* Ω[succ (succ n)] X) definition EM_up {G : AbGroup} {X : Type*} {n : }
: G →* Ω[succ n] (Ω X) := (e : AbInfGroup_of_AbGroup G →∞g Ωg[succ (succ n)] X) :
!loopn_succ_in ∘* e AbInfGroup_of_AbGroup G →∞g Ωg[succ n] (Ω X) :=
gloopn_succ_in (succ n) X ∘∞g e
definition is_homomorphism_EM_up {G : AbGroup} {X : Type*} {n : }
(e : G →* Ω[succ (succ n)] X)
(r : Π(g h : G), e (g * h) = e g ⬝ e h)
(g h : G) : EM_up e (g * h) = EM_up e g ⬝ EM_up e h :=
begin
refine ap !loopn_succ_in !r ⬝ _, apply apn_con,
end
definition EMadd1_pmap [unfold 8] {G : AbGroup} {X : Type*} (n : ) definition EMadd1_pmap [unfold 8] {G : AbGroup} {X : Type*} (n : )
(e : G →* Ω[succ n] X) (e : AbInfGroup_of_AbGroup G →∞g Ωg[succ n] X) [H : is_trunc (n.+1) X] : EMadd1 G n →* X :=
(r : Π(g h : G), e (g * h) = e g ⬝ e h)
[H : is_trunc (n.+1) X] : EMadd1 G n →* X :=
begin begin
revert X e r H, induction n with n f: intro X e r H, revert X e H, induction n with n f: intro X e H,
{ exact EM1_pmap e r }, { exact EM1_pmap e },
rewrite [EMadd1_succ], rewrite [EMadd1_succ],
apply ptrunc.elim ((succ n).+1), apply ptrunc.elim ((succ n).+1),
apply susp_elim, apply susp_elim,
exact f _ (EM_up e) (is_homomorphism_EM_up e r) _ exact f _ (EM_up e) _
end end
definition EMadd1_pmap_succ {G : AbGroup} {X : Type*} (n : ) (e : G →* Ω[succ (succ n)] X) definition EMadd1_pmap_succ {G : AbGroup} {X : Type*} (n : )
(r : Π(g h : G), e (g * h) = e g ⬝ e h) [H2 : is_trunc ((succ n).+1) X] : (e : AbInfGroup_of_AbGroup G →∞g Ωg[succ (succ n)] X) [H2 : is_trunc ((succ n).+1) X] :
EMadd1_pmap (succ n) e r = EMadd1_pmap (succ n) e = ptrunc.elim ((succ n).+1) (susp_elim (EMadd1_pmap n (EM_up e))) :=
ptrunc.elim ((succ n).+1) (susp_elim (EMadd1_pmap n (EM_up e) (is_homomorphism_EM_up e r))) :=
by reflexivity by reflexivity
definition loop_EMadd1_pmap {G : AbGroup} {X : Type*} {n : } (e : G →* Ω[succ (succ n)] X) definition loop_EMadd1_pmap {G : AbGroup} {X : Type*} {n : }
(r : Π(g h : G), e (g * h) = e g ⬝ e h) [H : is_trunc ((succ n).+1) X] : (e : AbInfGroup_of_AbGroup G →∞g Ωg[succ (succ n)] X) [H : is_trunc ((succ n).+1) X] :
Ω→(EMadd1_pmap (succ n) e r) ∘* loop_EMadd1 G n ~* Ω→(EMadd1_pmap (succ n) e) ∘* loop_EMadd1 G n ~* EMadd1_pmap n (EM_up e) :=
EMadd1_pmap n (EM_up e) (is_homomorphism_EM_up e r) :=
begin begin
cases n with n, cases n with n,
{ apply hopf_delooping_elim }, { apply hopf_delooping_elim },
@ -329,11 +315,11 @@ namespace EM
reflexivity } reflexivity }
end end
definition loopn_EMadd1_pmap' {G : AbGroup} {X : Type*} {n : } (e : G →* Ω[succ n] X) definition loopn_EMadd1_pmap' {G : AbGroup} {X : Type*} {n : }
(r : Π(g h : G), e (g * h) = e g ⬝ e h) [H : is_trunc (n.+1) X] : (e : AbInfGroup_of_AbGroup G →∞g Ωg[succ n] X) [H : is_trunc (n.+1) X] :
Ω→[succ n](EMadd1_pmap n e r) ∘* loopn_EMadd1 G n ~* e := Ω→[succ n](EMadd1_pmap n e) ∘* loopn_EMadd1 G n ~* pmap_of_inf_homomorphism e :=
begin begin
revert X e r H, induction n with n IH: intro X e r H, revert X e H, induction n with n IH: intro X e H,
{ apply loop_EM1_pmap }, { apply loop_EM1_pmap },
refine pwhisker_left _ !loopn_EMadd1_succ ⬝* _, refine pwhisker_left _ !loopn_EMadd1_succ ⬝* _,
refine !passoc⁻¹* ⬝* _, refine !passoc⁻¹* ⬝* _,
@ -341,14 +327,16 @@ namespace EM
refine !passoc ⬝* _, refine !passoc ⬝* _,
refine pwhisker_left _ (!passoc⁻¹* ⬝* refine pwhisker_left _ (!passoc⁻¹* ⬝*
pwhisker_right _ (!apn_pcompose⁻¹* ⬝* apn_phomotopy _ !loop_EMadd1_pmap) ⬝* !IH) ⬝* _, pwhisker_right _ (!apn_pcompose⁻¹* ⬝* apn_phomotopy _ !loop_EMadd1_pmap) ⬝* !IH) ⬝* _,
apply pinv_pcompose_cancel_left refine _ ⬝* pinv_pcompose_cancel_left !loopn_succ_in (pmap_of_inf_homomorphism e),
apply pwhisker_left,
apply phomotopy_of_homotopy, reflexivity, intro g, apply is_set_loopn,
end end
definition EMadd1_pequiv' {G : AbGroup} {X : Type*} (n : ) (e : G ≃* Ω[succ n] X) definition EMadd1_pequiv' {G : AbGroup} {X : Type*} (n : )
(r : Π(g h : G), e (g * h) = e g ⬝ e h) (e : AbInfGroup_of_AbGroup G ≃∞g Ωg[succ n] X) [H1 : is_conn n X] [H2 : is_trunc (n.+1) X] :
[H1 : is_conn n X] [H2 : is_trunc (n.+1) X] : EMadd1 G n ≃* X := EMadd1 G n ≃* X :=
begin begin
apply pequiv_of_pmap (EMadd1_pmap n e r), apply pequiv_of_pmap (EMadd1_pmap n e),
have is_conn 0 (EMadd1 G n), from is_conn_of_le _ (zero_le_of_nat n), have is_conn 0 (EMadd1 G n), from is_conn_of_le _ (zero_le_of_nat n),
have is_trunc (n.+1) (EMadd1 G n), from !is_trunc_EMadd1, have is_trunc (n.+1) (EMadd1 G n), from !is_trunc_EMadd1,
refine whitehead_principle_pointed (n.+1) _ _, refine whitehead_principle_pointed (n.+1) _ _,
@ -356,9 +344,9 @@ namespace EM
{ apply @is_equiv_of_is_contr, { apply @is_equiv_of_is_contr,
do 2 exact trivial_homotopy_group_of_is_conn _ (le_of_lt_succ H)}, do 2 exact trivial_homotopy_group_of_is_conn _ (le_of_lt_succ H)},
{ cases H, esimp, apply is_equiv_trunc_functor, esimp, { cases H, esimp, apply is_equiv_trunc_functor, esimp,
apply is_equiv.homotopy_closed, rotate 1, apply is_equiv.homotopy_closed,
{ symmetry, exact phomotopy_pinv_right_of_phomotopy (loopn_EMadd1_pmap' _ _) }, { symmetry, exact phomotopy_pinv_right_of_phomotopy (loopn_EMadd1_pmap' _) },
apply is_equiv_compose e, apply pequiv.to_is_equiv }, refine is_equiv_compose e _ _ _, apply inf_isomorphism.is_equiv_to_hom },
{ apply @is_equiv_of_is_contr, { apply @is_equiv_of_is_contr,
do 2 exact trivial_homotopy_group_of_is_trunc _ H} do 2 exact trivial_homotopy_group_of_is_trunc _ H}
end end
@ -366,13 +354,9 @@ namespace EM
definition EMadd1_pequiv {G : AbGroup} {X : Type*} (n : ) (e : G ≃g πg[n+1] X) definition EMadd1_pequiv {G : AbGroup} {X : Type*} (n : ) (e : G ≃g πg[n+1] X)
[H1 : is_conn n X] [H2 : is_trunc (n.+1) X] : EMadd1 G n ≃* X := [H1 : is_conn n X] [H2 : is_trunc (n.+1) X] : EMadd1 G n ≃* X :=
begin begin
have is_set (Ω[succ n] X), from !is_set_loopn, have is_set (Ωg[succ n] X), from is_set_loopn (succ n) X,
fapply EMadd1_pequiv' n, apply EMadd1_pequiv' n,
refine pequiv_of_isomorphism e ⬝e* !ptrunc_pequiv, refine inf_isomorphism_of_isomorphism e ⬝∞g gtrunc_isomorphism (Ωg[succ n] X),
refine preserve_binary_of_inv_preserve (pequiv_of_isomorphism e ⬝e* !ptrunc_pequiv) _ _ _,
-- apply EMadd1_pequiv' n ((ptrunc_pequiv _ _)⁻¹ᵉ* ⬝e* pequiv_of_isomorphism e⁻¹ᵍ),
-- apply inv_con
esimp, intro p q, exact to_respect_mul e⁻¹ᵍ (tr p) (tr q)
end end
definition EMadd1_pequiv_succ {G : AbGroup} {X : Type*} (n : ) (e : G ≃g πag[n+2] X) definition EMadd1_pequiv_succ {G : AbGroup} {X : Type*} (n : ) (e : G ≃g πag[n+2] X)

View file

@ -131,12 +131,12 @@ namespace chain_complex
(q : fiber_sequence_fun (n+1) (fiber.mk x p) = pt) (q : fiber_sequence_fun (n+1) (fiber.mk x p) = pt)
: fiber_sequence_carrier_pequiv n (fiber.mk (fiber.mk x p) q) : fiber_sequence_carrier_pequiv n (fiber.mk (fiber.mk x p) q)
= !respect_pt⁻¹ ⬝ ap (fiber_sequence_fun n) q⁻¹ ⬝ p := = !respect_pt⁻¹ ⬝ ap (fiber_sequence_fun n) q⁻¹ ⬝ p :=
fiber_ppoint_equiv_eq p q pfiber_ppoint_equiv_eq p q
definition fiber_sequence_carrier_pequiv_inv_eq (n : ) definition fiber_sequence_carrier_pequiv_inv_eq (n : )
(p : Ω(fiber_sequence_carrier n)) : (fiber_sequence_carrier_pequiv n)⁻¹ᵉ* p = (p : Ω(fiber_sequence_carrier n)) : (fiber_sequence_carrier_pequiv n)⁻¹ᵉ* p =
fiber.mk (fiber.mk pt (respect_pt (fiber_sequence_fun n) ⬝ p)) idp := fiber.mk (fiber.mk pt (respect_pt (fiber_sequence_fun n) ⬝ p)) idp :=
fiber_ppoint_equiv_inv_eq (fiber_sequence_fun n) p pfiber_ppoint_equiv_inv_eq (fiber_sequence_fun n) p
/- TODO: prove naturality of pfiber_ppoint_pequiv in general -/ /- TODO: prove naturality of pfiber_ppoint_pequiv in general -/

View file

@ -218,7 +218,7 @@ namespace is_conn
: is_surjective f → is_conn_fun -1 f := : is_surjective f → is_conn_fun -1 f :=
begin begin
intro H, intro b, intro H, intro b,
exact @is_contr_of_inhabited_prop (∥fiber f b∥) (is_trunc_trunc -1 (fiber f b)) (H b), exact is_contr_of_inhabited_prop (H b) _,
end end
definition is_surjection_of_minus_one_conn {A B : Type} (f : A → B) definition is_surjection_of_minus_one_conn {A B : Type} (f : A → B)
@ -232,7 +232,7 @@ namespace is_conn
λH, @center (∥A∥) H λH, @center (∥A∥) H
definition minus_one_conn_of_merely {A : Type} : ∥A∥ → is_conn -1 A := definition minus_one_conn_of_merely {A : Type} : ∥A∥ → is_conn -1 A :=
@is_contr_of_inhabited_prop (∥A∥) (is_trunc_trunc -1 A) λx, is_contr_of_inhabited_prop x _
section section
open arrow open arrow
@ -289,6 +289,9 @@ namespace is_conn
is_conn_fun k f := is_conn_fun k f :=
_ _
definition is_conn_fun_id (k : ℕ₋₂) (A : Type) : is_conn_fun k (@id A) :=
λa, _
-- Lemma 7.5.14 -- Lemma 7.5.14
theorem is_equiv_trunc_functor_of_is_conn_fun [instance] {A B : Type} (n : ℕ₋₂) (f : A → B) theorem is_equiv_trunc_functor_of_is_conn_fun [instance] {A B : Type} (n : ℕ₋₂) (f : A → B)
[H : is_conn_fun n f] : is_equiv (trunc_functor n f) := [H : is_conn_fun n f] : is_equiv (trunc_functor n f) :=
@ -308,7 +311,7 @@ namespace is_conn
[H2 : is_conn_fun k f] : is_conn_fun k (trunc_functor n f) := [H2 : is_conn_fun k f] : is_conn_fun k (trunc_functor n f) :=
begin begin
apply is_conn_fun.intro, apply is_conn_fun.intro,
intro P, have Πb, is_trunc n (P b), from (λb, is_trunc_of_le _ H), intro P, have Πb, is_trunc n (P b), from (λb, is_trunc_of_le _ H _),
fconstructor, fconstructor,
{ intro f' b, { intro f' b,
induction b with b, induction b with b,
@ -366,11 +369,11 @@ namespace is_conn
definition is_conn_succ_intro {n : ℕ₋₂} {A : Type} (a : trunc (n.+1) A) definition is_conn_succ_intro {n : ℕ₋₂} {A : Type} (a : trunc (n.+1) A)
(H2 : Π(a a' : A), is_conn n (a = a')) : is_conn (n.+1) A := (H2 : Π(a a' : A), is_conn n (a = a')) : is_conn (n.+1) A :=
begin begin
apply @is_contr_of_inhabited_prop, refine is_contr_of_inhabited_prop _ _,
{ exact a },
{ apply is_trunc_succ_intro, { apply is_trunc_succ_intro,
refine trunc.rec _, intro a, refine trunc.rec _, intro a', refine trunc.rec _, intro a, refine trunc.rec _, intro a',
exact is_contr_equiv_closed !tr_eq_tr_equiv⁻¹ᵉ _ }, exact is_contr_equiv_closed !tr_eq_tr_equiv⁻¹ᵉ _ }
exact a
end end
definition is_conn_pathover (n : ℕ₋₂) {A : Type} {B : A → Type} {a a' : A} (p : a = a') (b : B a) definition is_conn_pathover (n : ℕ₋₂) {A : Type} {B : A → Type} {a a' : A} (p : a = a') (b : B a)
@ -414,7 +417,7 @@ namespace is_conn
(H : k ≤ n) [H2 : is_conn_fun k f] : is_conn_fun k (trunc.elim f : trunc n A → B) := (H : k ≤ n) [H2 : is_conn_fun k f] : is_conn_fun k (trunc.elim f : trunc n A → B) :=
begin begin
apply is_conn_fun.intro, apply is_conn_fun.intro,
intro P, have Πb, is_trunc n (P b), from (λb, is_trunc_of_le _ H), intro P, have Πb, is_trunc n (P b), from (λb, is_trunc_of_le _ H _),
fconstructor, fconstructor,
{ intro f' b, { intro f' b,
refine is_conn_fun.elim k H2 _ _ b, intro a, exact f' (tr a) }, refine is_conn_fun.elim k H2 _ _ b, intro a, exact f' (tr a) },

View file

@ -49,9 +49,9 @@ namespace freudenthal section
definition is_equiv_code_merid (a : A) : is_equiv (code_merid a) := definition is_equiv_code_merid (a : A) : is_equiv (code_merid a) :=
begin begin
have Πa, is_trunc n.-2.+1 (is_equiv (code_merid a)), have Πa, is_trunc n.-2.+1 (is_equiv (code_merid a)),
from λa, is_trunc_of_le _ !minus_one_le_succ, from λa, is_trunc_of_le _ !minus_one_le_succ _,
refine is_conn.elim (n.-1) _ _ a, refine is_conn.elim (n.-1) _ _ a,
{ esimp, exact homotopy_closed id (homotopy.symm (code_merid_β_right))} { esimp, exact homotopy_closed id code_merid_β_right⁻¹ʰᵗʸ _ }
end end
definition code_merid_equiv [constructor] (a : A) : trunc (n + n) A ≃ trunc (n + n) A := definition code_merid_equiv [constructor] (a : A) : trunc (n + n) A ≃ trunc (n + n) A :=
@ -60,7 +60,7 @@ namespace freudenthal section
definition code_merid_inv_pt (x : trunc (n + n) A) : (code_merid_equiv pt)⁻¹ x = x := definition code_merid_inv_pt (x : trunc (n + n) A) : (code_merid_equiv pt)⁻¹ x = x :=
begin begin
refine ap010 @(is_equiv.inv _) _ x ⬝ _, refine ap010 @(is_equiv.inv _) _ x ⬝ _,
{ exact homotopy_closed id (homotopy.symm code_merid_β_right)}, { exact homotopy_closed id code_merid_β_right⁻¹ʰᵗʸ _ },
{ apply is_conn.elim_β}, { apply is_conn.elim_β},
{ reflexivity} { reflexivity}
end end

View file

@ -17,7 +17,7 @@ namespace is_trunc
begin begin
apply is_trunc_trunc_of_is_trunc, apply is_trunc_trunc_of_is_trunc,
apply is_contr_loop_of_is_trunc, apply is_contr_loop_of_is_trunc,
apply @is_trunc_of_le A n _, refine @is_trunc_of_le A n _ _ _,
apply trunc_index.le_of_succ_le_succ, apply trunc_index.le_of_succ_le_succ,
rewrite [succ_sub_two_succ k], rewrite [succ_sub_two_succ k],
exact of_nat_le_of_nat H, exact of_nat_le_of_nat H,
@ -84,7 +84,7 @@ namespace is_trunc
apply homotopy_group_functor_phomotopy, apply plift_functor_phomotopy apply homotopy_group_functor_phomotopy, apply plift_functor_phomotopy
end, end,
have π→[k] pdown.{v u} ∘ π→[k] (plift_functor f) ∘ π→[k] pup.{u v} ~ π→[k] f, from this, have π→[k] pdown.{v u} ∘ π→[k] (plift_functor f) ∘ π→[k] pup.{u v} ~ π→[k] f, from this,
apply is_equiv.homotopy_closed, rotate 1, apply is_equiv.homotopy_closed,
{ exact this}, { exact this},
{ do 2 apply is_equiv_compose, { do 2 apply is_equiv_compose,
{ apply is_equiv_homotopy_group_functor, apply to_is_equiv !equiv_lift}, { apply is_equiv_homotopy_group_functor, apply to_is_equiv !equiv_lift},
@ -112,7 +112,7 @@ namespace is_trunc
(H : Πa k, is_equiv (π→[k + 1] (pmap_of_map f a))) : is_equiv f := (H : Πa k, is_equiv (π→[k + 1] (pmap_of_map f a))) : is_equiv f :=
begin begin
revert A B HA HB f H' H, induction n with n IH: intros, revert A B HA HB f H' H, induction n with n IH: intros,
{ apply is_equiv_of_is_contr }, { exact is_equiv_of_is_contr _ _ _ },
have Πa, is_equiv (Ω→ (pmap_of_map f a)), have Πa, is_equiv (Ω→ (pmap_of_map f a)),
begin begin
intro a, intro a,
@ -124,7 +124,7 @@ namespace is_trunc
have Π(b : A) (p : a = b), have Π(b : A) (p : a = b),
is_equiv (pmap.to_fun (π→[k + 1] (pmap_of_map (ap f) p))), is_equiv (pmap.to_fun (π→[k + 1] (pmap_of_map (ap f) p))),
begin begin
intro b p, induction p, apply is_equiv.homotopy_closed, exact this, intro b p, induction p, refine is_equiv.homotopy_closed _ _ this,
refine homotopy_group_functor_phomotopy _ _, refine homotopy_group_functor_phomotopy _ _,
apply ap1_pmap_of_map apply ap1_pmap_of_map
end, end,
@ -135,7 +135,7 @@ namespace is_trunc
apply is_equiv_compose, exact this a p, apply is_equiv_compose, exact this a p,
exact is_equiv_trunc_functor _ _ _ exact is_equiv_trunc_functor _ _ _
end, end,
apply is_equiv.homotopy_closed, exact this, refine is_equiv.homotopy_closed _ _ this,
refine !homotopy_group_functor_pcompose⁻¹* ⬝* _, refine !homotopy_group_functor_pcompose⁻¹* ⬝* _,
apply homotopy_group_functor_phomotopy, apply homotopy_group_functor_phomotopy,
fapply phomotopy.mk, fapply phomotopy.mk,
@ -156,10 +156,10 @@ namespace is_trunc
begin begin
apply is_equiv_compose apply is_equiv_compose
(π→[k + 1] (pointed_eta_pequiv B ⬝e* (pequiv_of_eq_pt (respect_pt f))⁻¹ᵉ*)), (π→[k + 1] (pointed_eta_pequiv B ⬝e* (pequiv_of_eq_pt (respect_pt f))⁻¹ᵉ*)),
apply is_equiv_compose (π→[k + 1] f), refine is_equiv_compose (π→[k + 1] f) _ _ _,
all_goals exact is_equiv_homotopy_group_functor _ _ _, all_goals exact is_equiv_homotopy_group_functor _ _ _,
end, end,
refine @(is_equiv.homotopy_closed _) _ this _, refine is_equiv.homotopy_closed _ _ this,
apply to_homotopy, apply to_homotopy,
refine pwhisker_left _ !homotopy_group_functor_pcompose⁻¹* ⬝* _, refine pwhisker_left _ !homotopy_group_functor_pcompose⁻¹* ⬝* _,
refine !homotopy_group_functor_pcompose⁻¹* ⬝* _, refine !homotopy_group_functor_pcompose⁻¹* ⬝* _,
@ -170,12 +170,12 @@ namespace is_trunc
definition is_contr_of_trivial_homotopy (n : ℕ₋₂) (A : Type) [is_trunc n A] [is_conn 0 A] definition is_contr_of_trivial_homotopy (n : ℕ₋₂) (A : Type) [is_trunc n A] [is_conn 0 A]
(H : Πk a, is_contr (π[k] (pointed.MK A a))) : is_contr A := (H : Πk a, is_contr (π[k] (pointed.MK A a))) : is_contr A :=
begin begin
fapply is_trunc_is_equiv_closed_rev, { exact λa, ⋆}, refine is_trunc_is_equiv_closed_rev _ (λa, ⋆) _ _,
apply whitehead_principle n, apply whitehead_principle n,
{ apply is_equiv_trunc_functor_of_is_conn_fun, apply is_conn_fun_to_unit_of_is_conn}, { apply is_equiv_trunc_functor_of_is_conn_fun, apply is_conn_fun_to_unit_of_is_conn },
intro a k, intro a k,
apply @is_equiv_of_is_contr, refine is_equiv_of_is_contr _ _ _,
refine trivial_homotopy_group_of_is_trunc _ !zero_lt_succ, exact trivial_homotopy_group_of_is_trunc _ !zero_lt_succ,
end end
definition is_contr_of_trivial_homotopy_nat (n : ) (A : Type) [is_trunc n A] [is_conn 0 A] definition is_contr_of_trivial_homotopy_nat (n : ) (A : Type) [is_trunc n A] [is_conn 0 A]
@ -208,7 +208,7 @@ namespace is_trunc
definition is_trunc_of_trivial_homotopy {n : } {m : ℕ₋₂} {A : Type} (H : is_trunc m A) definition is_trunc_of_trivial_homotopy {n : } {m : ℕ₋₂} {A : Type} (H : is_trunc m A)
(H2 : Πk a, k > n → is_contr (π[k] (pointed.MK A a))) : is_trunc n A := (H2 : Πk a, k > n → is_contr (π[k] (pointed.MK A a))) : is_trunc n A :=
begin begin
fapply is_trunc_is_equiv_closed_rev, { exact @tr n A }, refine is_trunc_is_equiv_closed_rev _ (@tr n A) _ _,
apply whitehead_principle m, apply whitehead_principle m,
{ apply is_equiv_trunc_functor_of_is_conn_fun, { apply is_equiv_trunc_functor_of_is_conn_fun,
note z := is_conn_fun_tr n A, note z := is_conn_fun_tr n A,
@ -249,9 +249,9 @@ namespace is_trunc
begin begin
have is_conn 0 A, from !is_conn_of_is_conn_succ, have is_conn 0 A, from !is_conn_of_is_conn_succ,
cases n with n, cases n with n,
{ unfold [homotopy_group, ptrunc], apply ab_group_of_is_contr }, { unfold [homotopy_group, ptrunc], exact ab_group_of_is_contr _ _ },
cases n with n, cases n with n,
{ unfold [homotopy_group, ptrunc], apply ab_group_of_is_contr }, { unfold [homotopy_group, ptrunc], exact ab_group_of_is_contr _ _ },
exact ab_group_homotopy_group (n+2) A exact ab_group_homotopy_group (n+2) A
end end

View file

@ -37,7 +37,7 @@ section
begin begin
apply is_conn_fun.elim -1 (is_conn_fun_from_unit -1 A 1) apply is_conn_fun.elim -1 (is_conn_fun_from_unit -1 A 1)
(λa, trunctype.mk' -1 (is_equiv (λx, a * x))), (λa, trunctype.mk' -1 (is_equiv (λx, a * x))),
intro z, change is_equiv (λx : A, 1 * x), apply is_equiv.homotopy_closed id, intro z, change is_equiv (λx : A, 1 * x), refine is_equiv.homotopy_closed id _ _,
intro x, apply inverse, apply one_mul intro x, apply inverse, apply one_mul
end end
@ -45,7 +45,7 @@ section
begin begin
apply is_conn_fun.elim -1 (is_conn_fun_from_unit -1 A 1) apply is_conn_fun.elim -1 (is_conn_fun_from_unit -1 A 1)
(λa, trunctype.mk' -1 (is_equiv (λx, x * a))), (λa, trunctype.mk' -1 (is_equiv (λx, x * a))),
intro z, change is_equiv (λx : A, x * 1), apply is_equiv.homotopy_closed id, intro z, change is_equiv (λx : A, x * 1), refine is_equiv.homotopy_closed id _ _,
intro x, apply inverse, apply mul_one intro x, apply inverse, apply mul_one
end end
end end

View file

@ -45,10 +45,10 @@ namespace is_equiv
is_equiv.mk id id (λa, idp) (λa, idp) (λa, idp) is_equiv.mk id id (λa, idp) (λa, idp) (λa, idp)
-- The composition of two equivalences is, again, an equivalence. -- The composition of two equivalences is, again, an equivalence.
definition is_equiv_compose [constructor] [Hf : is_equiv f] [Hg : is_equiv g] definition is_equiv_compose [constructor] (Hf : is_equiv f) (Hg : is_equiv g)
: is_equiv (g ∘ f) := : is_equiv (g ∘ f) :=
is_equiv.mk (g ∘ f) (f⁻¹ ∘ g⁻¹) is_equiv.mk (g ∘ f) (f⁻¹ ∘ g⁻¹)
abstract (λc, ap g (right_inv f (g⁻¹ c)) ⬝ right_inv g c) end abstract (λc, ap g (right_inv f (g⁻¹ c)) ⬝ right_inv g c) end
abstract (λa, ap (inv f) (left_inv g (f a)) ⬝ left_inv f a) end abstract (λa, ap (inv f) (left_inv g (f a)) ⬝ left_inv f a) end
abstract (λa, (whisker_left _ (adj g (f a))) ⬝ abstract (λa, (whisker_left _ (adj g (f a))) ⬝
(ap_con g _ _)⁻¹ ⬝ (ap_con g _ _)⁻¹ ⬝
@ -105,23 +105,23 @@ namespace is_equiv
end end
-- Any function pointwise equal to an equivalence is an equivalence as well. -- Any function pointwise equal to an equivalence is an equivalence as well.
definition homotopy_closed [constructor] {A B : Type} (f : A → B) {f' : A → B} [Hf : is_equiv f] definition homotopy_closed [constructor] {A B : Type} (f : A → B) {f' : A → B} (Hty : f ~ f')
(Hty : f ~ f') : is_equiv f' := (Hf : is_equiv f) : is_equiv f' :=
adjointify f' adjointify f'
(inv f) (inv f)
(λ b, (Hty (inv f b))⁻¹ ⬝ right_inv f b) (λ b, (Hty (inv f b))⁻¹ ⬝ right_inv f b)
(λ a, (ap (inv f) (Hty a))⁻¹ ⬝ left_inv f a) (λ a, (ap (inv f) (Hty a))⁻¹ ⬝ left_inv f a)
definition inv_homotopy_closed [constructor] {A B : Type} {f : A → B} {f' : B → A} definition inv_homotopy_closed [constructor] {A B : Type} {f : A → B} (f' : B → A)
[Hf : is_equiv f] (Hty : f⁻¹ ~ f') : is_equiv f := (Hf : is_equiv f) (Hty : f⁻¹ᶠ ~ f') : is_equiv f :=
adjointify f adjointify f
f' f'
(λ b, ap f !Hty⁻¹ ⬝ right_inv f b) (λ b, ap f !Hty⁻¹ ⬝ right_inv f b)
(λ a, !Hty⁻¹ ⬝ left_inv f a) (λ a, !Hty⁻¹ ⬝ left_inv f a)
definition inv_homotopy_inv {A B : Type} {f g : A → B} [is_equiv f] [is_equiv g] (p : f ~ g) definition inv_homotopy_inv {A B : Type} {f g : A → B} [is_equiv f] [is_equiv g] (p : f ~ g)
: f⁻¹ ~ g⁻¹ := : f⁻¹ ~ g⁻¹ :=
λb, (left_inv g (f⁻¹ b))⁻¹ ⬝ ap g⁻¹ ((p (f⁻¹ b))⁻¹ ⬝ right_inv f b) λb, (left_inv g (f⁻¹ b))⁻¹ ⬝ ap g⁻¹ ((p (f⁻¹ b))⁻¹ ⬝ right_inv f b)
definition is_equiv_up [instance] [constructor] (A : Type) definition is_equiv_up [instance] [constructor] (A : Type)
: is_equiv (up : A → lift A) := : is_equiv (up : A → lift A) :=
@ -140,47 +140,45 @@ namespace is_equiv
-- over all of B. -- over all of B.
definition is_equiv_rect (P : B → Type) (g : Πa, P (f a)) (b : B) : P b := definition is_equiv_rect (P : B → Type) (g : Πa, P (f a)) (b : B) : P b :=
right_inv f b ▸ g (f⁻¹ b) right_inv f b ▸ g (f⁻¹ b)
definition is_equiv_rect' (P : A → B → Type) (g : Πb, P (f⁻¹ b) b) (a : A) : P a (f a) := definition is_equiv_rect' (P : A → B → Type) (g : Πb, P (f⁻¹ b) b) (a : A) : P a (f a) :=
left_inv f a ▸ g (f a) left_inv f a ▸ g (f a)
definition is_equiv_rect_comp (P : B → Type) definition is_equiv_rect_comp (P : B → Type)
(df : Π (x : A), P (f x)) (x : A) : is_equiv_rect f P df (f x) = df x := (df : Π (x : A), P (f x)) (x : A) : is_equiv_rect f P df (f x) = df x :=
calc calc
is_equiv_rect f P df (f x) is_equiv_rect f P df (f x)
= right_inv f (f x) ▸ df (f⁻¹ (f x)) : by esimp = right_inv f (f x) ▸ df (f⁻¹ (f x)) : by esimp
... = ap f (left_inv f x) ▸ df (f⁻¹ (f x)) : by rewrite -adj ... = ap f (left_inv f x) ▸ df (f⁻¹ (f x)) : by rewrite -adj
... = left_inv f x ▸ df (f⁻¹ (f x)) : by rewrite -tr_compose ... = left_inv f x ▸ df (f⁻¹ (f x)) : by rewrite -tr_compose
... = df x : by rewrite (apdt df (left_inv f x)) ... = df x : by rewrite (apdt df (left_inv f x))
theorem adj_inv (b : B) : left_inv f (f⁻¹ b) = ap f⁻¹ (right_inv f b) := theorem adj_inv (b : B) : left_inv f (f⁻¹ b) = ap f⁻¹ (right_inv f b) :=
is_equiv_rect f _ is_equiv_rect f _
(λa, eq.cancel_right (left_inv f (id a)) (λa, eq.cancel_right (left_inv f (id a))
(whisker_left _ !ap_id⁻¹ ⬝ (ap_con_eq_con_ap (left_inv f) (left_inv f a))⁻¹) ⬝ (whisker_left _ !ap_id⁻¹ ⬝ (ap_con_eq_con_ap (left_inv f) (left_inv f a))⁻¹) ⬝
!ap_compose ⬝ ap02 f⁻¹ (adj f a)⁻¹) !ap_compose ⬝ ap02 f⁻¹ (adj f a)⁻¹)
b b
--The inverse of an equivalence is, again, an equivalence. --The inverse of an equivalence is, again, an equivalence.
definition is_equiv_inv [instance] [constructor] [priority 500] : is_equiv f⁻¹ := definition is_equiv_inv [instance] [constructor] [priority 500] : is_equiv f⁻¹ :=
is_equiv.mk f⁻¹ f (left_inv f) (right_inv f) (adj_inv f) is_equiv.mk f⁻¹ f (left_inv f) (right_inv f) (adj_inv f)
-- The 2-out-of-3 properties -- The 2-out-of-3 properties
definition cancel_right (g : B → C) [Hgf : is_equiv (g ∘ f)] : (is_equiv g) := definition cancel_right (g : B → C) [Hgf : is_equiv (g ∘ f)] : (is_equiv g) :=
have Hfinv : is_equiv f⁻¹, from is_equiv_inv f, homotopy_closed _ (λb, ap g (right_inv f b)) (is_equiv_compose (g ∘ f) f⁻¹ᶠ _ _)
@homotopy_closed _ _ _ _ (is_equiv_compose (g ∘ f) f⁻¹) (λb, ap g (@right_inv _ _ f _ b))
definition cancel_left (g : C → A) [Hgf : is_equiv (f ∘ g)] : (is_equiv g) := definition cancel_left (g : C → A) [Hgf : is_equiv (f ∘ g)] : (is_equiv g) :=
have Hfinv : is_equiv f⁻¹, from is_equiv_inv f, homotopy_closed _ (λa, left_inv f (g a)) (is_equiv_compose f⁻¹ᶠ (f ∘ g) _ _)
@homotopy_closed _ _ _ _ (is_equiv_compose f⁻¹ (f ∘ g)) (λa, left_inv f (g a))
definition inj' [unfold 4] {x y : A} (q : f x = f y) : x = y := definition inj' [unfold 4] {x y : A} (q : f x = f y) : x = y :=
(left_inv f x)⁻¹ ⬝ ap f⁻¹ q ⬝ left_inv f y (left_inv f x)⁻¹ ⬝ ap f⁻¹ q ⬝ left_inv f y
definition ap_inj' {x y : A} (q : f x = f y) : ap f (inj' f q) = q := definition ap_inj' {x y : A} (q : f x = f y) : ap f (inj' f q) = q :=
!ap_con ⬝ whisker_right _ !ap_con !ap_con ⬝ whisker_right _ !ap_con
⬝ ((!ap_inv ⬝ inverse2 (adj f _)⁻¹) ⬝ ((!ap_inv ⬝ inverse2 (adj f _)⁻¹)
◾ (inverse (ap_compose f f⁻¹ _)) ◾ (inverse (ap_compose f f⁻¹ _))
◾ (adj f _)⁻¹) ◾ (adj f _)⁻¹)
⬝ con_ap_con_eq_con_con (right_inv f) _ _ ⬝ con_ap_con_eq_con_con (right_inv f) _ _
⬝ whisker_right _ !con.left_inv ⬝ whisker_right _ !con.left_inv
@ -204,16 +202,16 @@ namespace is_equiv
section rewrite_rules section rewrite_rules
variables {a : A} {b : B} variables {a : A} {b : B}
definition eq_of_eq_inv (p : a = f⁻¹ b) : f a = b := definition eq_of_eq_inv (p : a = f⁻¹ b) : f a = b :=
ap f p ⬝ right_inv f b ap f p ⬝ right_inv f b
definition eq_of_inv_eq (p : f⁻¹ b = a) : b = f a := definition eq_of_inv_eq (p : f⁻¹ b = a) : b = f a :=
(eq_of_eq_inv p⁻¹)⁻¹ (eq_of_eq_inv p⁻¹)⁻¹
definition inv_eq_of_eq (p : b = f a) : f⁻¹ b = a := definition inv_eq_of_eq (p : b = f a) : f⁻¹ b = a :=
ap f⁻¹ p ⬝ left_inv f a ap f⁻¹ p ⬝ left_inv f a
definition eq_inv_of_eq (p : f a = b) : a = f⁻¹ b := definition eq_inv_of_eq (p : f a = b) : a = f⁻¹ b :=
(inv_eq_of_eq p⁻¹)⁻¹ (inv_eq_of_eq p⁻¹)⁻¹
end rewrite_rules end rewrite_rules
@ -222,33 +220,33 @@ namespace is_equiv
section pre_compose section pre_compose
variables (α : A → C) (β : B → C) variables (α : A → C) (β : B → C)
definition homotopy_of_homotopy_inv_pre (p : β ~ α ∘ f⁻¹) : β ∘ f ~ α := definition homotopy_of_homotopy_inv_pre (p : β ~ α ∘ f⁻¹) : β ∘ f ~ α :=
λ a, p (f a) ⬝ ap α (left_inv f a) λ a, p (f a) ⬝ ap α (left_inv f a)
definition homotopy_of_inv_homotopy_pre (p : α ∘ f⁻¹ ~ β) : α ~ β ∘ f := definition homotopy_of_inv_homotopy_pre (p : α ∘ f⁻¹ ~ β) : α ~ β ∘ f :=
λ a, (ap α (left_inv f a))⁻¹ ⬝ p (f a) λ a, (ap α (left_inv f a))⁻¹ ⬝ p (f a)
definition inv_homotopy_of_homotopy_pre (p : α ~ β ∘ f) : α ∘ f⁻¹ ~ β := definition inv_homotopy_of_homotopy_pre (p : α ~ β ∘ f) : α ∘ f⁻¹ ~ β :=
λ b, p (f⁻¹ b) ⬝ ap β (right_inv f b) λ b, p (f⁻¹ b) ⬝ ap β (right_inv f b)
definition homotopy_inv_of_homotopy_pre (p : β ∘ f ~ α) : β ~ α ∘ f⁻¹ := definition homotopy_inv_of_homotopy_pre (p : β ∘ f ~ α) : β ~ α ∘ f⁻¹ :=
λ b, (ap β (right_inv f b))⁻¹ ⬝ p (f⁻¹ b) λ b, (ap β (right_inv f b))⁻¹ ⬝ p (f⁻¹ b)
end pre_compose end pre_compose
section post_compose section post_compose
variables (α : C → A) (β : C → B) variables (α : C → A) (β : C → B)
definition homotopy_of_homotopy_inv_post (p : α ~ f⁻¹ ∘ β) : f ∘ α ~ β := definition homotopy_of_homotopy_inv_post (p : α ~ f⁻¹ ∘ β) : f ∘ α ~ β :=
λ c, ap f (p c) ⬝ (right_inv f (β c)) λ c, ap f (p c) ⬝ (right_inv f (β c))
definition homotopy_of_inv_homotopy_post (p : f⁻¹ ∘ β ~ α) : β ~ f ∘ α := definition homotopy_of_inv_homotopy_post (p : f⁻¹ ∘ β ~ α) : β ~ f ∘ α :=
λ c, (right_inv f (β c))⁻¹ ⬝ ap f (p c) λ c, (right_inv f (β c))⁻¹ ⬝ ap f (p c)
definition inv_homotopy_of_homotopy_post (p : β ~ f ∘ α) : f⁻¹ ∘ β ~ α := definition inv_homotopy_of_homotopy_post (p : β ~ f ∘ α) : f⁻¹ ∘ β ~ α :=
λ c, ap f⁻¹ (p c) ⬝ (left_inv f (α c)) λ c, ap f⁻¹ (p c) ⬝ (left_inv f (α c))
definition homotopy_inv_of_homotopy_post (p : f ∘ α ~ β) : α ~ f⁻¹ ∘ β := definition homotopy_inv_of_homotopy_post (p : f ∘ α ~ β) : α ~ f⁻¹ ∘ β :=
λ c, (left_inv f (α c))⁻¹ ⬝ ap f⁻¹ (p c) λ c, (left_inv f (α c))⁻¹ ⬝ ap f⁻¹ (p c)
end post_compose end post_compose
end end
@ -268,16 +266,16 @@ namespace is_equiv
include H include H
definition inv_commute' (p : Π⦃a : A⦄ (b : B (g' a)), f (h b) = h' (f b)) {a : A} definition inv_commute' (p : Π⦃a : A⦄ (b : B (g' a)), f (h b) = h' (f b)) {a : A}
(c : C (g' a)) : f⁻¹ (h' c) = h (f⁻¹ c) := (c : C (g' a)) : f⁻¹ (h' c) = h (f⁻¹ c) :=
inj' f (right_inv f (h' c) ⬝ ap h' (right_inv f c)⁻¹ ⬝ (p (f⁻¹ c))⁻¹) inj' f (right_inv f (h' c) ⬝ ap h' (right_inv f c)⁻¹ ⬝ (p (f⁻¹ c))⁻¹)
definition fun_commute_of_inv_commute' (p : Π⦃a : A⦄ (c : C (g' a)), f⁻¹ (h' c) = h (f⁻¹ c)) definition fun_commute_of_inv_commute' (p : Π⦃a : A⦄ (c : C (g' a)), f⁻¹ (h' c) = h (f⁻¹ c))
{a : A} (b : B (g' a)) : f (h b) = h' (f b) := {a : A} (b : B (g' a)) : f (h b) = h' (f b) :=
inj' f⁻¹ (left_inv f (h b) ⬝ ap h (left_inv f b)⁻¹ ⬝ (p (f b))⁻¹) inj' f⁻¹ (left_inv f (h b) ⬝ ap h (left_inv f b)⁻¹ ⬝ (p (f b))⁻¹)
definition ap_inv_commute' (p : Π⦃a : A⦄ (b : B (g' a)), f (h b) = h' (f b)) {a : A} definition ap_inv_commute' (p : Π⦃a : A⦄ (b : B (g' a)), f (h b) = h' (f b)) {a : A}
(c : C (g' a)) : ap f (inv_commute' @f @h @h' p c) (c : C (g' a)) : ap f (inv_commute' @f @h @h' p c)
= right_inv f (h' c) ⬝ ap h' (right_inv f c)⁻¹ ⬝ (p (f⁻¹ c))⁻¹ := = right_inv f (h' c) ⬝ ap h' (right_inv f c)⁻¹ ⬝ (p (f⁻¹ c))⁻¹ :=
!ap_inj' !ap_inj'
-- inv_commute'_fn is in types.equiv -- inv_commute'_fn is in types.equiv
@ -285,8 +283,8 @@ namespace is_equiv
-- This is inv_commute' for A ≡ unit -- This is inv_commute' for A ≡ unit
definition inv_commute1' {B C : Type} (f : B → C) [is_equiv f] (h : B → B) (h' : C → C) definition inv_commute1' {B C : Type} (f : B → C) [is_equiv f] (h : B → B) (h' : C → C)
(p : Π(b : B), f (h b) = h' (f b)) (c : C) : f⁻¹ (h' c) = h (f⁻¹ c) := (p : Π(b : B), f (h b) = h' (f b)) (c : C) : f⁻¹ (h' c) = h (f⁻¹ c) :=
inj' f (right_inv f (h' c) ⬝ ap h' (right_inv f c)⁻¹ ⬝ (p (f⁻¹ c))⁻¹) inj' f (right_inv f (h' c) ⬝ ap h' (right_inv f c)⁻¹ ⬝ (p (f⁻¹ c))⁻¹)
end is_equiv end is_equiv
open is_equiv open is_equiv
@ -316,10 +314,10 @@ namespace equiv
(right_inv : Πb, f (g b) = b) (left_inv : Πa, g (f a) = a) : A ≃ B := (right_inv : Πb, f (g b) = b) (left_inv : Πa, g (f a) = a) : A ≃ B :=
equiv.mk f (adjointify f g right_inv left_inv) equiv.mk f (adjointify f g right_inv left_inv)
definition to_inv [reducible] [unfold 3] (f : A ≃ B) : B → A := f⁻¹ definition to_inv [reducible] [unfold 3] (f : A ≃ B) : B → A := f⁻¹
definition to_right_inv [reducible] [unfold 3] (f : A ≃ B) (b : B) : f (f⁻¹ b) = b := definition to_right_inv [reducible] [unfold 3] (f : A ≃ B) (b : B) : f (f⁻¹ b) = b :=
right_inv f b right_inv f b
definition to_left_inv [reducible] [unfold 3] (f : A ≃ B) (a : A) : f⁻¹ (f a) = a := definition to_left_inv [reducible] [unfold 3] (f : A ≃ B) (a : A) : f⁻¹ (f a) = a :=
left_inv f a left_inv f a
protected definition rfl [refl] [constructor] : A ≃ A := protected definition rfl [refl] [constructor] : A ≃ A :=
@ -329,10 +327,10 @@ namespace equiv
@equiv.rfl A @equiv.rfl A
protected definition symm [symm] [constructor] (f : A ≃ B) : B ≃ A := protected definition symm [symm] [constructor] (f : A ≃ B) : B ≃ A :=
equiv.mk f⁻¹ !is_equiv_inv equiv.mk f⁻¹ !is_equiv_inv
protected definition trans [trans] [constructor] (f : A ≃ B) (g : B ≃ C) : A ≃ C := protected definition trans [trans] [constructor] (f : A ≃ B) (g : B ≃ C) : A ≃ C :=
equiv.mk (g ∘ f) !is_equiv_compose equiv.mk (g ∘ f) (is_equiv_compose g f _ _)
infixl ` ⬝e `:75 := equiv.trans infixl ` ⬝e `:75 := equiv.trans
postfix `⁻¹ᵉ`:(max + 1) := equiv.symm postfix `⁻¹ᵉ`:(max + 1) := equiv.symm
@ -344,11 +342,11 @@ namespace equiv
idp idp
definition equiv_change_fun [constructor] (f : A ≃ B) {f' : A → B} (Heq : f ~ f') : A ≃ B := definition equiv_change_fun [constructor] (f : A ≃ B) {f' : A → B} (Heq : f ~ f') : A ≃ B :=
equiv.mk f' (is_equiv.homotopy_closed f Heq) equiv.mk f' (is_equiv.homotopy_closed f Heq _)
definition equiv_change_inv [constructor] (f : A ≃ B) {f' : B → A} (Heq : f⁻¹ ~ f') definition equiv_change_inv [constructor] (f : A ≃ B) {f' : B → A} (Heq : f⁻¹ ~ f')
: A ≃ B := : A ≃ B :=
equiv.mk f (inv_homotopy_closed Heq) equiv.mk f (inv_homotopy_closed _ _ Heq)
definition eq_equiv_fn_eq_of_is_equiv [constructor] (f : A → B) [H : is_equiv f] (a b : A) : definition eq_equiv_fn_eq_of_is_equiv [constructor] (f : A → B) [H : is_equiv f] (a b : A) :
(a = b) ≃ (f a = f b) := (a = b) ≃ (f a = f b) :=
@ -368,9 +366,9 @@ namespace equiv
idp idp
definition inj [unfold 3] (f : A ≃ B) {x y : A} (q : f x = f y) : x = y := definition inj [unfold 3] (f : A ≃ B) {x y : A} (q : f x = f y) : x = y :=
(left_inv f x)⁻¹ ⬝ ap f⁻¹ q ⬝ left_inv f y (left_inv f x)⁻¹ ⬝ ap f⁻¹ q ⬝ left_inv f y
definition inj_inv [unfold 3] (f : A ≃ B) {x y : B} (q : f⁻¹ x = f⁻¹ y) : x = y := definition inj_inv [unfold 3] (f : A ≃ B) {x y : B} (q : f⁻¹ x = f⁻¹ y) : x = y :=
(right_inv f x)⁻¹ ⬝ ap f q ⬝ right_inv f y (right_inv f x)⁻¹ ⬝ ap f q ⬝ right_inv f y
definition ap_inj (f : A ≃ B) {x y : A} (q : f x = f y) : ap f (inj' f q) = q := definition ap_inj (f : A ≃ B) {x y : A} (q : f x = f y) : ap f (inj' f q) = q :=
@ -394,34 +392,34 @@ namespace equiv
definition equiv_lift [constructor] (A : Type) : A ≃ lift A := equiv.mk up _ definition equiv_lift [constructor] (A : Type) : A ≃ lift A := equiv.mk up _
definition equiv_rect (f : A ≃ B) (P : B → Type) (g : Πa, P (f a)) (b : B) : P b := definition equiv_rect (f : A ≃ B) (P : B → Type) (g : Πa, P (f a)) (b : B) : P b :=
right_inv f b ▸ g (f⁻¹ b) right_inv f b ▸ g (f⁻¹ b)
definition equiv_rect' (f : A ≃ B) (P : A → B → Type) (g : Πb, P (f⁻¹ b) b) (a : A) : P a (f a) := definition equiv_rect' (f : A ≃ B) (P : A → B → Type) (g : Πb, P (f⁻¹ b) b) (a : A) : P a (f a) :=
left_inv f a ▸ g (f a) left_inv f a ▸ g (f a)
definition equiv_rect_comp (f : A ≃ B) (P : B → Type) definition equiv_rect_comp (f : A ≃ B) (P : B → Type)
(df : Π (x : A), P (f x)) (x : A) : equiv_rect f P df (f x) = df x := (df : Π (x : A), P (f x)) (x : A) : equiv_rect f P df (f x) = df x :=
calc calc
equiv_rect f P df (f x) equiv_rect f P df (f x)
= right_inv f (f x) ▸ df (f⁻¹ (f x)) : by esimp = right_inv f (f x) ▸ df (f⁻¹ (f x)) : by esimp
... = ap f (left_inv f x) ▸ df (f⁻¹ (f x)) : by rewrite -adj ... = ap f (left_inv f x) ▸ df (f⁻¹ (f x)) : by rewrite -adj
... = left_inv f x ▸ df (f⁻¹ (f x)) : by rewrite -tr_compose ... = left_inv f x ▸ df (f⁻¹ (f x)) : by rewrite -tr_compose
... = df x : by rewrite (apdt df (left_inv f x)) ... = df x : by rewrite (apdt df (left_inv f x))
end end
section section
variables {A B : Type} (f : A ≃ B) {a : A} {b : B} variables {A B : Type} (f : A ≃ B) {a : A} {b : B}
definition to_eq_of_eq_inv (p : a = f⁻¹ b) : f a = b := definition to_eq_of_eq_inv (p : a = f⁻¹ b) : f a = b :=
ap f p ⬝ right_inv f b ap f p ⬝ right_inv f b
definition to_eq_of_inv_eq (p : f⁻¹ b = a) : b = f a := definition to_eq_of_inv_eq (p : f⁻¹ b = a) : b = f a :=
(eq_of_eq_inv p⁻¹)⁻¹ (eq_of_eq_inv p⁻¹)⁻¹
definition to_inv_eq_of_eq (p : b = f a) : f⁻¹ b = a := definition to_inv_eq_of_eq (p : b = f a) : f⁻¹ b = a :=
ap f⁻¹ p ⬝ left_inv f a ap f⁻¹ p ⬝ left_inv f a
definition to_eq_inv_of_eq (p : f a = b) : a = f⁻¹ b := definition to_eq_inv_of_eq (p : f a = b) : a = f⁻¹ b :=
(inv_eq_of_eq p⁻¹)⁻¹ (inv_eq_of_eq p⁻¹)⁻¹
end end
@ -432,15 +430,15 @@ namespace equiv
{g : A → A} {g' : A → A} (h : Π{a}, B (g' a) → B (g a)) (h' : Π{a}, C (g' a) → C (g a)) {g : A → A} {g' : A → A} (h : Π{a}, B (g' a) → B (g a)) (h' : Π{a}, C (g' a) → C (g a))
definition inv_commute (p : Π⦃a : A⦄ (b : B (g' a)), f (h b) = h' (f b)) {a : A} definition inv_commute (p : Π⦃a : A⦄ (b : B (g' a)), f (h b) = h' (f b)) {a : A}
(c : C (g' a)) : f⁻¹ (h' c) = h (f⁻¹ c) := (c : C (g' a)) : f⁻¹ (h' c) = h (f⁻¹ c) :=
inv_commute' @f @h @h' p c inv_commute' @f @h @h' p c
definition fun_commute_of_inv_commute (p : Π⦃a : A⦄ (c : C (g' a)), f⁻¹ (h' c) = h (f⁻¹ c)) definition fun_commute_of_inv_commute (p : Π⦃a : A⦄ (c : C (g' a)), f⁻¹ (h' c) = h (f⁻¹ c))
{a : A} (b : B (g' a)) : f (h b) = h' (f b) := {a : A} (b : B (g' a)) : f (h b) = h' (f b) :=
fun_commute_of_inv_commute' @f @h @h' p b fun_commute_of_inv_commute' @f @h @h' p b
definition inv_commute1 {B C : Type} (f : B ≃ C) (h : B → B) (h' : C → C) definition inv_commute1 {B C : Type} (f : B ≃ C) (h : B → B) (h' : C → C)
(p : Π(b : B), f (h b) = h' (f b)) (c : C) : f⁻¹ (h' c) = h (f⁻¹ c) := (p : Π(b : B), f (h b) = h' (f b)) (c : C) : f⁻¹ (h' c) = h (f⁻¹ c) :=
inv_commute1' (to_fun f) h h' p c inv_commute1' (to_fun f) h h' p c
end end
@ -455,7 +453,7 @@ namespace is_equiv
definition is_equiv_of_equiv_of_homotopy [constructor] {A B : Type} (f : A ≃ B) definition is_equiv_of_equiv_of_homotopy [constructor] {A B : Type} (f : A ≃ B)
{f' : A → B} (Hty : f ~ f') : is_equiv f' := {f' : A → B} (Hty : f ~ f') : is_equiv f' :=
@(homotopy_closed f) f' _ Hty homotopy_closed f Hty _
end is_equiv end is_equiv

View file

@ -13,7 +13,7 @@ open equiv is_equiv function
variables {A A' : Type} {B B' : A → Type} {B'' : A' → Type} {C : Π⦃a⦄, B a → Type} variables {A A' : Type} {B B' : A → Type} {B'' : A' → Type} {C : Π⦃a⦄, B a → Type}
{a a₂ a₃ a₄ : A} {p p' : a = a₂} {p₂ : a₂ = a₃} {p₃ : a₃ = a₄} {p₁₃ : a = a₃} {a a₂ a₃ a₄ : A} {p p' : a = a₂} {p₂ : a₂ = a₃} {p₃ : a₃ = a₄} {p₁₃ : a = a₃}
{b b' : B a} {b₂ b₂' : B a₂} {b₃ : B a₃} {b₄ : B a₄} {a' a₂' a₃' : A'} {b b' : B a} {b₂ b₂' : B a₂} {b₃ : B a₃} {b₄ : B a₄}
{c : C b} {c₂ : C b₂} {c : C b} {c₂ : C b₂}
namespace eq namespace eq
@ -127,10 +127,10 @@ namespace eq
definition cono.left_inv (r : b =[p] b₂) : r⁻¹ᵒ ⬝o r =[!con.left_inv] idpo := definition cono.left_inv (r : b =[p] b₂) : r⁻¹ᵒ ⬝o r =[!con.left_inv] idpo :=
by induction r; constructor by induction r; constructor
definition eq_of_pathover {a' a₂' : A'} (q : a' =[p] a₂') : a' = a₂' := definition eq_of_pathover (q : a' =[p] a₂') : a' = a₂' :=
by cases q;reflexivity by cases q;reflexivity
definition pathover_of_eq [unfold 5 8] (p : a = a₂) {a' a₂' : A'} (q : a' = a₂') : a' =[p] a₂' := definition pathover_of_eq [unfold 5 8] (p : a = a₂) (q : a' = a₂') : a' =[p] a₂' :=
by cases p;cases q;constructor by cases p;cases q;constructor
definition pathover_constant [constructor] (p : a = a₂) (a' a₂' : A') : a' =[p] a₂' ≃ a' = a₂' := definition pathover_constant [constructor] (p : a = a₂) (a' a₂' : A') : a' =[p] a₂' ≃ a' = a₂' :=
@ -159,8 +159,8 @@ namespace eq
(to_right_inv !pathover_equiv_tr_eq) (to_right_inv !pathover_equiv_tr_eq)
(to_left_inv !pathover_equiv_tr_eq) (to_left_inv !pathover_equiv_tr_eq)
definition eq_of_pathover_idp_pathover_of_eq {A X : Type} (x : X) {a a' : A} (p : a = a') : definition eq_of_pathover_idp_pathover_of_eq (a' : A') (p : a = a₂) :
eq_of_pathover_idp (pathover_of_eq (idpath x) p) = p := eq_of_pathover_idp (pathover_of_eq (idpath a') p) = p :=
by induction p; reflexivity by induction p; reflexivity
variable (B) variable (B)
@ -327,6 +327,7 @@ namespace eq
definition apd_inv (f : Πa, B a) (p : a = a₂) : apd f p⁻¹ = (apd f p)⁻¹ᵒ := definition apd_inv (f : Πa, B a) (p : a = a₂) : apd f p⁻¹ = (apd f p)⁻¹ᵒ :=
by cases p; reflexivity by cases p; reflexivity
/- probably more useful: apd_eq_ap -/
definition apd_eq_pathover_of_eq_ap (f : A → A') (p : a = a₂) : definition apd_eq_pathover_of_eq_ap (f : A → A') (p : a = a₂) :
apd f p = pathover_of_eq p (ap f p) := apd f p = pathover_of_eq p (ap f p) :=
eq.rec_on p idp eq.rec_on p idp
@ -466,4 +467,32 @@ namespace eq
apd0111 f (ap k p) (pathover_ap B k (apo l q)) (pathover_tro _ (m c)) := apd0111 f (ap k p) (pathover_ap B k (apo l q)) (pathover_tro _ (m c)) :=
by induction q; reflexivity by induction q; reflexivity
/- some properties about eq_of_pathover -/
definition apd_eq_ap (f : A → A') (p : a = a₂) : eq_of_pathover (apd f p) = ap f p :=
begin induction p, reflexivity end
definition eq_of_pathover_idp_constant (p : a' =[idpath a] a₂') :
eq_of_pathover_idp p = eq_of_pathover p :=
begin induction p using idp_rec_on, reflexivity end
definition eq_of_pathover_change_path (r : p = p') (q : a' =[p] a₂') :
eq_of_pathover (change_path r q) = eq_of_pathover q :=
begin induction r, reflexivity end
definition eq_of_pathover_cono (q : a' =[p] a₂') (q₂ : a₂' =[p₂] a₃') :
eq_of_pathover (q ⬝o q₂) = eq_of_pathover q ⬝ eq_of_pathover q₂ :=
begin induction q₂, reflexivity end
definition eq_of_pathover_invo (q : a' =[p] a₂') :
eq_of_pathover q⁻¹ᵒ = (eq_of_pathover q)⁻¹ :=
begin induction q, reflexivity end
definition eq_of_pathover_concato_eq (q : a' =[p] a₂') (q₂ : a₂' = a₃') :
eq_of_pathover (q ⬝op q₂) = eq_of_pathover q ⬝ q₂ :=
begin induction q₂, reflexivity end
definition eq_of_pathover_eq_concato (q : a' = a₂') (q₂ : a₂' =[p₂] a₃') :
eq_of_pathover (q ⬝po q₂) = q ⬝ eq_of_pathover q₂ :=
begin induction q, induction q₂, reflexivity end
end eq end eq

View file

@ -173,7 +173,7 @@ namespace is_trunc
--in the proof the type of H is given explicitly to make it available for class inference --in the proof the type of H is given explicitly to make it available for class inference
theorem is_trunc_of_le.{l} (A : Type.{l}) {n m : ℕ₋₂} (Hnm : n ≤ m) theorem is_trunc_of_le.{l} (A : Type.{l}) {n m : ℕ₋₂} (Hnm : n ≤ m)
[Hn : is_trunc n A] : is_trunc m A := (Hn : is_trunc n A) : is_trunc m A :=
begin begin
induction Hnm with m Hnm IH, induction Hnm with m Hnm IH,
{ exact Hn}, { exact Hn},
@ -191,23 +191,21 @@ namespace is_trunc
end end
-- these must be definitions, because we need them to compute sometimes -- these must be definitions, because we need them to compute sometimes
definition is_trunc_of_is_contr (A : Type) (n : ℕ₋₂) [H : is_contr A] : is_trunc n A := definition is_trunc_of_is_contr (A : Type) (n : ℕ₋₂) (H : is_contr A) : is_trunc n A :=
trunc_index.rec_on n H (λn H, _) trunc_index.rec_on n H (λn H, _)
definition is_trunc_succ_of_is_prop (A : Type) (n : ℕ₋₂) [H : is_prop A] definition is_trunc_succ_of_is_prop (A : Type) (n : ℕ₋₂) (H : is_prop A) : is_trunc (n.+1) A :=
: is_trunc (n.+1) A := is_trunc_of_le A (show -1 ≤ n.+1, from succ_le_succ (minus_two_le n)) _
is_trunc_of_le A (show -1 ≤ n.+1, from succ_le_succ (minus_two_le n))
definition is_trunc_succ_succ_of_is_set (A : Type) (n : ℕ₋₂) [H : is_set A] definition is_trunc_succ_succ_of_is_set (A : Type) (n : ℕ₋₂) (H : is_set A) : is_trunc (n.+2) A :=
: is_trunc (n.+2) A := is_trunc_of_le A (show 0 ≤ n.+2, from succ_le_succ (succ_le_succ (minus_two_le n))) _
is_trunc_of_le A (show 0 ≤ n.+2, from succ_le_succ (succ_le_succ (minus_two_le n)))
/- props -/ /- propositions -/
definition is_prop.elim [H : is_prop A] (x y : A) : x = y := definition is_prop.elim [H : is_prop A] (x y : A) : x = y :=
!center !center
definition is_contr_of_inhabited_prop {A : Type} [H : is_prop A] (x : A) : is_contr A := definition is_contr_of_inhabited_prop {A : Type} (x : A) (H : is_prop A) : is_contr A :=
is_contr.mk x (λy, !is_prop.elim) is_contr.mk x (λy, !is_prop.elim)
theorem is_prop_of_imp_is_contr {A : Type} (H : A → is_contr A) : is_prop A := theorem is_prop_of_imp_is_contr {A : Type} (H : A → is_contr A) : is_prop A :=
@ -257,64 +255,68 @@ namespace is_trunc
local attribute is_contr_unit is_prop_empty [instance] local attribute is_contr_unit is_prop_empty [instance]
definition is_trunc_unit [instance] (n : ℕ₋₂) : is_trunc n unit := definition is_trunc_unit [instance] (n : ℕ₋₂) : is_trunc n unit :=
!is_trunc_of_is_contr is_trunc_of_is_contr _ _ _
definition is_trunc_empty [instance] (n : ℕ₋₂) : is_trunc (n.+1) empty := definition is_trunc_empty [instance] (n : ℕ₋₂) : is_trunc (n.+1) empty :=
!is_trunc_succ_of_is_prop is_trunc_succ_of_is_prop _ _ _
/- interaction with equivalences -/ /- interaction with equivalences -/
section section
open is_equiv equiv open is_equiv equiv
definition is_contr_is_equiv_closed (f : A → B) [Hf : is_equiv f] [HA: is_contr A] definition is_contr_is_equiv_closed (f : A → B) (Hf : is_equiv f) (HA: is_contr A)
: (is_contr B) := : (is_contr B) :=
is_contr.mk (f (center A)) (λp, eq_of_eq_inv !center_eq) is_contr.mk (f (center A)) (λp, eq_of_eq_inv !center_eq)
definition is_contr_equiv_closed (H : A ≃ B) (HA : is_contr A) : is_contr B := definition is_contr_equiv_closed (H : A ≃ B) (HA : is_contr A) : is_contr B :=
is_contr_is_equiv_closed (to_fun H) is_contr_is_equiv_closed (to_fun H) _ _
definition is_contr_equiv_closed_rev (H : A ≃ B) (HB : is_contr B) : is_contr A := definition is_contr_equiv_closed_rev (H : A ≃ B) (HB : is_contr B) : is_contr A :=
is_contr_equiv_closed H⁻¹ᵉ HB is_contr_equiv_closed H⁻¹ᵉ HB
definition equiv_of_is_contr_of_is_contr [HA : is_contr A] [HB : is_contr B] : A ≃ B := definition equiv_of_is_contr_of_is_contr (HA : is_contr A) (HB : is_contr B) : A ≃ B :=
equiv.mk equiv.mk
(λa, center B) (λa, center B)
(is_equiv.adjointify (λa, center B) (λb, center A) center_eq center_eq) (is_equiv.adjointify (λa, center B) (λb, center A) center_eq center_eq)
theorem is_trunc_is_equiv_closed (n : ℕ₋₂) (f : A → B) [H : is_equiv f] theorem is_trunc_is_equiv_closed (n : ℕ₋₂) (f : A → B) (H : is_equiv f)
[HA : is_trunc n A] : is_trunc n B := (HA : is_trunc n A) : is_trunc n B :=
begin begin
revert A HA B f H, induction n with n IH: intros, revert A HA B f H, induction n with n IH: intros,
{ exact is_contr_is_equiv_closed f}, { exact is_contr_is_equiv_closed f _ _ },
{ apply is_trunc_succ_intro, intro x y, { apply is_trunc_succ_intro, intro x y,
exact IH (f⁻¹ x = f⁻¹ y) _ (x = y) (ap f⁻¹)⁻¹ !is_equiv_inv} exact IH (f⁻¹ x = f⁻¹ y) _ (x = y) (ap f⁻¹)⁻¹ !is_equiv_inv }
end end
definition is_trunc_is_equiv_closed_rev (n : ℕ₋₂) (f : A → B) [H : is_equiv f] definition is_trunc_is_equiv_closed_rev (n : ℕ₋₂) (f : A → B) (H : is_equiv f)
[HA : is_trunc n B] : is_trunc n A := (HA : is_trunc n B) : is_trunc n A :=
is_trunc_is_equiv_closed n f⁻¹ is_trunc_is_equiv_closed n f⁻¹ᶠ _ _
definition is_trunc_equiv_closed (n : ℕ₋₂) (f : A ≃ B) (HA : is_trunc n A) : is_trunc n B := definition is_trunc_equiv_closed (n : ℕ₋₂) (f : A ≃ B) (HA : is_trunc n A) : is_trunc n B :=
is_trunc_is_equiv_closed n (to_fun f) is_trunc_is_equiv_closed n (to_fun f) _ _
definition is_trunc_equiv_closed_rev (n : ℕ₋₂) (f : A ≃ B) (HA : is_trunc n B) : is_trunc n A := definition is_trunc_equiv_closed_rev (n : ℕ₋₂) (f : A ≃ B) (HA : is_trunc n B) : is_trunc n A :=
is_trunc_is_equiv_closed n (to_inv f) is_trunc_is_equiv_closed n (to_inv f) _ _
definition is_equiv_of_is_prop [constructor] [HA : is_prop A] [HB : is_prop B] definition is_equiv_of_is_prop [constructor] (f : A → B) (g : B → A)
(f : A → B) (g : B → A) : is_equiv f := (HA : is_prop A) (HB : is_prop B) : is_equiv f :=
is_equiv.mk f g (λb, !is_prop.elim) (λa, !is_prop.elim) (λa, !is_set.elim) is_equiv.mk f g (λb, !is_prop.elim) (λa, !is_prop.elim) (λa, !is_set.elim)
definition is_equiv_of_is_contr [constructor] [HA : is_contr A] [HB : is_contr B] definition is_equiv_of_is_contr [constructor] (f : A → B)
(f : A → B) : is_equiv f := (HA : is_contr A) (HB : is_contr B) : is_equiv f :=
is_equiv.mk f (λx, !center) (λb, !is_prop.elim) (λa, !is_prop.elim) (λa, !is_set.elim) is_equiv.mk f (λx, !center) (λb, !is_prop.elim) (λa, !is_prop.elim) (λa, !is_set.elim)
definition equiv_of_is_prop [constructor] [HA : is_prop A] [HB : is_prop B] definition equiv_of_is_contr [constructor] (HA : is_contr A) (HB : is_contr B) : A ≃ B :=
(f : A → B) (g : B → A) : A ≃ B := equiv.mk (λa, center B) (is_equiv_of_is_contr _ _ _)
equiv.mk f (is_equiv_of_is_prop f g)
definition equiv_of_iff_of_is_prop [unfold 5] [HA : is_prop A] [HB : is_prop B] (H : A ↔ B) : A ≃ B := definition equiv_of_is_prop [constructor] (f : A → B) (g : B → A)
equiv_of_is_prop (iff.elim_left H) (iff.elim_right H) (HA : is_prop A) (HB : is_prop B) : A ≃ B :=
equiv.mk f (is_equiv_of_is_prop f g _ _)
definition equiv_of_iff_of_is_prop [unfold 5] (HA : is_prop A) (HB : is_prop B) (H : A ↔ B) :
A ≃ B :=
equiv_of_is_prop (iff.elim_left H) (iff.elim_right H) _ _
/- truncatedness of lift -/ /- truncatedness of lift -/
definition is_trunc_lift [instance] [priority 1450] (A : Type) (n : ℕ₋₂) definition is_trunc_lift [instance] [priority 1450] (A : Type) (n : ℕ₋₂)
@ -328,11 +330,8 @@ namespace is_trunc
open equiv open equiv
/- A contractible type is equivalent to unit. -/ /- A contractible type is equivalent to unit. -/
variable (A) variable (A)
definition equiv_unit_of_is_contr [constructor] [H : is_contr A] : A ≃ unit := definition equiv_unit_of_is_contr [constructor] (H : is_contr A) : A ≃ unit :=
equiv.MK (λ (x : A), ⋆) equiv_of_is_contr _ _
(λ (u : unit), center A)
(λ (u : unit), unit.rec_on u idp)
(λ (x : A), center_eq x)
/- interaction with pathovers -/ /- interaction with pathovers -/
variable {A} variable {A}

View file

@ -40,20 +40,16 @@ namespace is_trunc
begin begin
induction n, induction n,
{ intro A, { intro A,
apply is_trunc_is_equiv_closed, apply is_trunc_equiv_closed _ !is_contr.sigma_char,
{ apply equiv.to_is_equiv, apply is_contr.sigma_char},
apply is_prop.mk, intros, apply is_prop.mk, intros,
fapply sigma_eq, apply x.2, fapply sigma_eq, apply x.2,
apply is_prop.elimo}, apply is_prop.elimo },
{ intro A, { intro A, exact is_trunc_equiv_closed _ !is_trunc.pi_char _ },
apply is_trunc_is_equiv_closed,
apply equiv.to_is_equiv,
apply is_trunc.pi_char},
end end
local attribute is_prop_is_trunc [instance] local attribute is_prop_is_trunc [instance]
definition is_trunc_succ_is_trunc [instance] (n m : ℕ₋₂) (A : Type) : definition is_trunc_succ_is_trunc [instance] (n m : ℕ₋₂) (A : Type) :
is_trunc (n.+1) (is_trunc m A) := is_trunc (n.+1) (is_trunc m A) :=
!is_trunc_succ_of_is_prop is_trunc_succ_of_is_prop _ _ _
end is_trunc end is_trunc

View file

@ -98,7 +98,7 @@ namespace is_equiv
@is_equiv_of_is_contr_fun _ _ f (λb, @is_contr_fiber_of_is_equiv _ _ _ (H b) _) @is_equiv_of_is_contr_fun _ _ f (λb, @is_contr_fiber_of_is_equiv _ _ _ (H b) _)
definition is_equiv_equiv_is_contr_fun : is_equiv f ≃ is_contr_fun f := definition is_equiv_equiv_is_contr_fun : is_equiv f ≃ is_contr_fun f :=
equiv_of_is_prop _ (λH, !is_equiv_of_is_contr_fun) equiv_of_is_prop _ (λH, !is_equiv_of_is_contr_fun) _ _
theorem inv_commute'_fn {A : Type} {B C : A → Type} (f : Π{a}, B a → C a) [H : Πa, is_equiv (@f a)] theorem inv_commute'_fn {A : Type} {B C : A → Type} (f : Π{a}, B a → C a) [H : Πa, is_equiv (@f a)]
{g : A → A} (h : Π{a}, B a → B (g a)) (h' : Π{a}, C a → C (g a)) {g : A → A} (h : Π{a}, B a → B (g a)) (h' : Π{a}, C a → C (g a))
@ -300,16 +300,17 @@ namespace equiv
definition is_contr_equiv (A B : Type) [HA : is_contr A] [HB : is_contr B] : is_contr (A ≃ B) := definition is_contr_equiv (A B : Type) [HA : is_contr A] [HB : is_contr B] : is_contr (A ≃ B) :=
begin begin
apply @is_contr_of_inhabited_prop, apply is_prop.mk, refine is_contr_of_inhabited_prop _ _,
intro x y, cases x with fx Hx, cases y with fy Hy, generalize Hy, { exact equiv_of_is_contr_of_is_contr _ _ },
apply (eq_of_homotopy (λ a, !eq_of_is_contr)) ▸ (λ Hy, !is_prop.elim ▸ rfl), { apply is_prop.mk,
apply equiv_of_is_contr_of_is_contr intro x y, cases x with fx Hx, cases y with fy Hy, generalize Hy,
apply (eq_of_homotopy (λ a, !eq_of_is_contr)) ▸ (λ Hy, !is_prop.elim ▸ rfl) }
end end
definition is_trunc_succ_equiv (n : trunc_index) (A B : Type) definition is_trunc_succ_equiv (n : trunc_index) (A B : Type)
[HA : is_trunc n.+1 A] [HB : is_trunc n.+1 B] : is_trunc n.+1 (A ≃ B) := [HA : is_trunc n.+1 A] [HB : is_trunc n.+1 B] : is_trunc n.+1 (A ≃ B) :=
@is_trunc_equiv_closed _ _ n.+1 (equiv.symm !equiv.sigma_char) @is_trunc_equiv_closed _ _ n.+1 (equiv.symm !equiv.sigma_char)
(@is_trunc_sigma _ _ _ _ (λ f, !is_trunc_succ_of_is_prop)) (@is_trunc_sigma _ _ _ _ (λ f, is_trunc_succ_of_is_prop _ _ _))
definition is_trunc_equiv (n : trunc_index) (A B : Type) definition is_trunc_equiv (n : trunc_index) (A B : Type)
[HA : is_trunc n A] [HB : is_trunc n B] : is_trunc n (A ≃ B) := [HA : is_trunc n A] [HB : is_trunc n B] : is_trunc n (A ≃ B) :=

View file

@ -8,14 +8,15 @@ Theorems about fibers
-/ -/
import .sigma .eq .pi cubical.squareover .pointed .eq import .sigma .eq .pi cubical.squareover .pointed .eq
open equiv sigma sigma.ops eq pi pointed is_equiv open equiv sigma sigma.ops eq pi pointed is_equiv is_trunc function unit
structure fiber {A B : Type} (f : A → B) (b : B) := structure fiber {A B : Type} (f : A → B) (b : B) :=
(point : A) (point : A)
(point_eq : f point = b) (point_eq : f point = b)
variables {A B : Type} {f : A → B} {b : B}
namespace fiber namespace fiber
variables {A B : Type} {f : A → B} {b : B}
protected definition sigma_char [constructor] protected definition sigma_char [constructor]
(f : A → B) (b : B) : fiber f b ≃ (Σ(a : A), f a = b) := (f : A → B) (b : B) : fiber f b ≃ (Σ(a : A), f a = b) :=
@ -27,7 +28,8 @@ namespace fiber
{intro x, cases x, apply idp }, {intro x, cases x, apply idp },
end end
definition fiber_eq_equiv [constructor] (x y : fiber f b) /- equality type of a fiber -/
definition fiber_eq_equiv' [constructor] (x y : fiber f b)
: (x = y) ≃ (Σ(p : point x = point y), point_eq x = ap f p ⬝ point_eq y) := : (x = y) ≃ (Σ(p : point x = point y), point_eq x = ap f p ⬝ point_eq y) :=
begin begin
apply equiv.trans, apply equiv.trans,
@ -41,7 +43,77 @@ namespace fiber
definition fiber_eq {x y : fiber f b} (p : point x = point y) definition fiber_eq {x y : fiber f b} (p : point x = point y)
(q : point_eq x = ap f p ⬝ point_eq y) : x = y := (q : point_eq x = ap f p ⬝ point_eq y) : x = y :=
to_inv !fiber_eq_equiv ⟨p, q⟩ to_inv !fiber_eq_equiv' ⟨p, q⟩
definition fiber_eq_equiv [constructor] (x y : fiber f b) :
(x = y) ≃ (Σ(p : point x = point y), point_eq x = ap f p ⬝ point_eq y) :=
@equiv_change_inv _ _ (fiber_eq_equiv' x y) (λpq, fiber_eq pq.1 pq.2)
begin intro pq, cases pq, reflexivity end
definition point_fiber_eq {x y : fiber f b}
(p : point x = point y) (q : point_eq x = ap f p ⬝ point_eq y) :
ap point (fiber_eq p q) = p :=
begin
induction x with a r, induction y with a' s, esimp at *, induction p,
induction q using eq.rec_symm, induction s, reflexivity
end
definition fiber_eq_equiv_fiber (x y : fiber f b) :
x = y ≃ fiber (ap1_gen f (point_eq x) (point_eq y)) (idpath b) :=
calc
x = y ≃ fiber.sigma_char f b x = fiber.sigma_char f b y :
eq_equiv_fn_eq (fiber.sigma_char f b) x y
... ≃ Σ(p : point x = point y), point_eq x =[p] point_eq y : sigma_eq_equiv
... ≃ Σ(p : point x = point y), (point_eq x)⁻¹ ⬝ ap f p ⬝ point_eq y = idp :
sigma_equiv_sigma_right (λp,
calc point_eq x =[p] point_eq y ≃ point_eq x = ap f p ⬝ point_eq y : eq_pathover_equiv_Fl
... ≃ ap f p ⬝ point_eq y = point_eq x : eq_equiv_eq_symm
... ≃ (point_eq x)⁻¹ ⬝ (ap f p ⬝ point_eq y) = idp : eq_equiv_inv_con_eq_idp
... ≃ (point_eq x)⁻¹ ⬝ ap f p ⬝ point_eq y = idp : equiv_eq_closed_left _ !con.assoc⁻¹)
... ≃ fiber (ap1_gen f (point_eq x) (point_eq y)) (idpath b) : fiber.sigma_char
definition point_fiber_eq_equiv_fiber {x y : fiber f b}
(p : x = y) : point (fiber_eq_equiv_fiber x y p) = ap1_gen point idp idp p :=
by induction p; reflexivity
definition fiber_eq_pr2 {x y : fiber f b}
(p : x = y) : point_eq x = ap f (ap point p) ⬝ point_eq y :=
begin induction p, exact !idp_con⁻¹ end
definition fiber_eq_eta {x y : fiber f b}
(p : x = y) : p = fiber_eq (ap point p) (fiber_eq_pr2 p) :=
begin induction p, induction x with a q, induction q, reflexivity end
definition fiber_eq_con {x y z : fiber f b}
(p1 : point x = point y) (p2 : point y = point z)
(q1 : point_eq x = ap f p1 ⬝ point_eq y) (q2 : point_eq y = ap f p2 ⬝ point_eq z) :
fiber_eq p1 q1 ⬝ fiber_eq p2 q2 =
fiber_eq (p1 ⬝ p2) (q1 ⬝ whisker_left (ap f p1) q2 ⬝ !con.assoc⁻¹ ⬝
whisker_right (point_eq z) (ap_con f p1 p2)⁻¹) :=
begin
induction x with a₁ r₁, induction y with a₂ r₂, induction z with a₃ r₃, esimp at *,
induction q2 using eq.rec_symm, induction q1 using eq.rec_symm,
induction p2, induction p1, induction r₃, reflexivity
end
definition fiber_eq2_equiv {x y : fiber f b}
(p₁ p₂ : point x = point y) (q₁ : point_eq x = ap f p₁ ⬝ point_eq y)
(q₂ : point_eq x = ap f p₂ ⬝ point_eq y) : (fiber_eq p₁ q₁ = fiber_eq p₂ q₂) ≃
(Σ(r : p₁ = p₂), q₁ ⬝ whisker_right (point_eq y) (ap02 f r) = q₂) :=
begin
refine (eq_equiv_fn_eq (fiber_eq_equiv x y)⁻¹ᵉ _ _)⁻¹ᵉ ⬝e sigma_eq_equiv _ _ ⬝e _,
apply sigma_equiv_sigma_right, esimp, intro r,
refine !eq_pathover_equiv_square ⬝e _,
refine eq_hconcat_equiv !ap_constant ⬝e hconcat_eq_equiv (ap_compose (λx, x ⬝ _) _ _) ⬝e _,
refine !square_equiv_eq ⬝e _,
exact eq_equiv_eq_closed idp (idp_con q₂)
end
definition fiber_eq2 {x y : fiber f b}
{p₁ p₂ : point x = point y} {q₁ : point_eq x = ap f p₁ ⬝ point_eq y}
{q₂ : point_eq x = ap f p₂ ⬝ point_eq y} (r : p₁ = p₂)
(s : q₁ ⬝ whisker_right (point_eq y) (ap02 f r) = q₂) : (fiber_eq p₁ q₁ = fiber_eq p₂ q₂) :=
(fiber_eq2_equiv p₁ p₂ q₁ q₂)⁻¹ᵉ ⟨r, s⟩
definition fiber_pathover {X : Type} {A B : X → Type} {x₁ x₂ : X} {p : x₁ = x₂} definition fiber_pathover {X : Type} {A B : X → Type} {x₁ x₂ : X} {p : x₁ = x₂}
{f : Πx, A x → B x} {b : Πx, B x} {v₁ : fiber (f x₁) (b x₁)} {v₂ : fiber (f x₂) (b x₂)} {f : Πx, A x → B x} {b : Πx, B x} {v₁ : fiber (f x₁) (b x₁)} {v₂ : fiber (f x₂) (b x₂)}
@ -57,7 +129,38 @@ namespace fiber
apply pathover_idp_of_eq, apply eq_of_vdeg_square, apply square_of_squareover_ids r} apply pathover_idp_of_eq, apply eq_of_vdeg_square, apply square_of_squareover_ids r}
end end
open is_trunc definition is_trunc_fiber (n : ℕ₋₂) {A B : Type} (f : A → B) (b : B)
[is_trunc n A] [is_trunc (n.+1) B] : is_trunc n (fiber f b) :=
is_trunc_equiv_closed_rev n !fiber.sigma_char _
definition is_contr_fiber_id (A : Type) (a : A) : is_contr (fiber (@id A) a) :=
is_contr.mk (fiber.mk a idp)
begin intro x, induction x with a p, esimp at p, cases p, reflexivity end
/- the general functoriality between fibers -/
-- todo: show that this is an equivalence if g and h are, and use that for the special cases below
definition fiber_functor [constructor] {A A' B B' : Type} {f : A → B} {f' : A' → B'}
{b : B} {b' : B'} (g : A → A') (h : B → B') (H : hsquare g h f f') (p : h b = b')
(x : fiber f b) : fiber f' b' :=
fiber.mk (g (point x)) (H (point x) ⬝ ap h (point_eq x) ⬝ p)
/- equivalences between fibers -/
definition fiber_equiv_of_homotopy {A B : Type} {f g : A → B} (h : f ~ g) (b : B)
: fiber f b ≃ fiber g b :=
begin
refine (fiber.sigma_char f b ⬝e _ ⬝e (fiber.sigma_char g b)⁻¹ᵉ),
apply sigma_equiv_sigma_right, intros a,
apply equiv_eq_closed_left, apply h
end
definition fiber_equiv_basepoint [constructor] {A B : Type} (f : A → B) {b1 b2 : B} (p : b1 = b2)
: fiber f b1 ≃ fiber f b2 :=
calc fiber f b1 ≃ Σa, f a = b1 : fiber.sigma_char
... ≃ Σa, f a = b2 : sigma_equiv_sigma_right (λa, equiv_eq_closed_right (f a) p)
... ≃ fiber f b2 : fiber.sigma_char
definition fiber_pr1 (B : A → Type) (a : A) : fiber (pr1 : (Σa, B a) → A) a ≃ B a := definition fiber_pr1 (B : A → Type) (a : A) : fiber (pr1 : (Σa, B a) → A) a ≃ B a :=
calc calc
fiber pr1 a ≃ Σu, u.1 = a : fiber.sigma_char fiber pr1 a ≃ Σu, u.1 = a : fiber.sigma_char
@ -72,20 +175,7 @@ namespace fiber
... ≃ Σa b, f a = b : sigma_comm_equiv ... ≃ Σa b, f a = b : sigma_comm_equiv
... ≃ A : sigma_equiv_of_is_contr_right ... ≃ A : sigma_equiv_of_is_contr_right
definition is_pointed_fiber [instance] [constructor] (f : A → B) (a : A)
: pointed (fiber f (f a)) :=
pointed.mk (fiber.mk a idp)
definition pointed_fiber [constructor] (f : A → B) (a : A) : Type* :=
pointed.Mk (fiber.mk a (idpath (f a)))
definition is_trunc_fun [reducible] (n : ℕ₋₂) (f : A → B) :=
Π(b : B), is_trunc n (fiber f b)
definition is_contr_fun [reducible] (f : A → B) := is_trunc_fun -2 f
-- pre and post composition with equivalences -- pre and post composition with equivalences
open function
variable (f) variable (f)
protected definition equiv_postcompose [constructor] {B' : Type} (g : B ≃ B') --[H : is_equiv g] protected definition equiv_postcompose [constructor] {B' : Type} (g : B ≃ B') --[H : is_equiv g]
(b : B) : fiber (g ∘ f) (g b) ≃ fiber f b := (b : B) : fiber (g ∘ f) (g b) ≃ fiber f b :=
@ -107,11 +197,16 @@ namespace fiber
end end
... ≃ fiber f b : fiber.sigma_char ... ≃ fiber f b : fiber.sigma_char
end fiber definition fiber_equiv_of_square {A B C D : Type} {b : B} {d : D} {f : A → B} {g : C → D}
(h : A ≃ C) (k : B ≃ D) (s : k ∘ f ~ g ∘ h) (p : k b = d) : fiber f b ≃ fiber g d :=
calc fiber f b ≃ fiber (k ∘ f) (k b) : fiber.equiv_postcompose
... ≃ fiber (k ∘ f) d : fiber_equiv_basepoint (k ∘ f) p
... ≃ fiber (g ∘ h) d : fiber_equiv_of_homotopy s d
... ≃ fiber g d : fiber.equiv_precompose
open unit is_trunc pointed definition fiber_equiv_of_triangle {A B C : Type} {b : B} {f : A → B} {g : C → B} (h : A ≃ C)
(s : f ~ g ∘ h) : fiber f b ≃ fiber g b :=
namespace fiber fiber_equiv_of_square h erfl s idp
definition fiber_star_equiv [constructor] (A : Type) : fiber (λx : A, star) star ≃ A := definition fiber_star_equiv [constructor] (A : Type) : fiber (λx : A, star) star ≃ A :=
begin begin
@ -130,8 +225,48 @@ namespace fiber
≃ Σz : unit, a₀ = a : fiber.sigma_char ≃ Σz : unit, a₀ = a : fiber.sigma_char
... ≃ a₀ = a : sigma_unit_left ... ≃ a₀ = a : sigma_unit_left
-- the pointed fiber of a pointed map, which is the fiber over the basepoint definition fiber_total_equiv [constructor] {P Q : A → Type} (f : Πa, P a → Q a) {a : A} (q : Q a)
open pointed : fiber (total f) ⟨a , q⟩ ≃ fiber (f a) q :=
calc
fiber (total f) ⟨a , q⟩
≃ Σ(w : Σx, P x), ⟨w.1 , f w.1 w.2 ⟩ = ⟨a , q⟩
: fiber.sigma_char
... ≃ Σ(x : A), Σ(p : P x), ⟨x , f x p⟩ = ⟨a , q⟩
: sigma_assoc_equiv
... ≃ Σ(x : A), Σ(p : P x), Σ(H : x = a), f x p =[H] q
:
begin
apply sigma_equiv_sigma_right, intro x,
apply sigma_equiv_sigma_right, intro p,
apply sigma_eq_equiv
end
... ≃ Σ(x : A), Σ(H : x = a), Σ(p : P x), f x p =[H] q
:
begin
apply sigma_equiv_sigma_right, intro x,
apply sigma_comm_equiv
end
... ≃ Σ(w : Σx, x = a), Σ(p : P w.1), f w.1 p =[w.2] q
: sigma_assoc_equiv
... ≃ Σ(p : P (center (Σx, x=a)).1), f (center (Σx, x=a)).1 p =[(center (Σx, x=a)).2] q
: sigma_equiv_of_is_contr_left
... ≃ Σ(p : P a), f a p =[idpath a] q
: equiv_of_eq idp
... ≃ Σ(p : P a), f a p = q
:
begin
apply sigma_equiv_sigma_right, intro p,
apply pathover_idp
end
... ≃ fiber (f a) q
: fiber.sigma_char
definition fiber_equiv_of_is_contr [constructor] {A B : Type} (f : A → B) (b : B) [is_contr B] :
fiber f b ≃ A :=
!fiber.sigma_char ⬝e !sigma_equiv_of_is_contr_right
/- the pointed fiber of a pointed map, which is the fiber over the basepoint -/
definition pfiber [constructor] {X Y : Type*} (f : X →* Y) : Type* := definition pfiber [constructor] {X Y : Type*} (f : X →* Y) : Type* :=
pointed.MK (fiber f pt) (fiber.mk pt !respect_pt) pointed.MK (fiber f pt) (fiber.mk pt !respect_pt)
@ -150,25 +285,17 @@ namespace fiber
: pfiber f ≃* pfiber g := : pfiber f ≃* pfiber g :=
begin begin
fapply pequiv_of_equiv, fapply pequiv_of_equiv,
{ refine (fiber.sigma_char f pt ⬝e _ ⬝e (fiber.sigma_char g pt)⁻¹ᵉ), { exact fiber_equiv_of_homotopy h pt },
apply sigma_equiv_sigma_right, intros a,
apply equiv_eq_closed_left, apply (to_homotopy h) },
{ refine (fiber_eq rfl _), { refine (fiber_eq rfl _),
change (h pt)⁻¹ ⬝ respect_pt f = idp ⬝ respect_pt g, change (h pt)⁻¹ ⬝ respect_pt f = idp ⬝ respect_pt g,
rewrite idp_con, apply inv_con_eq_of_eq_con, symmetry, exact (to_homotopy_pt h) } rewrite idp_con, apply inv_con_eq_of_eq_con, symmetry, exact (to_homotopy_pt h) }
end end
definition transport_fiber_equiv [constructor] {A B : Type} (f : A → B) {b1 b2 : B} (p : b1 = b2)
: fiber f b1 ≃ fiber f b2 :=
calc fiber f b1 ≃ Σa, f a = b1 : fiber.sigma_char
... ≃ Σa, f a = b2 : sigma_equiv_sigma_right (λa, equiv_eq_closed_right (f a) p)
... ≃ fiber f b2 : fiber.sigma_char
definition pequiv_postcompose {A B B' : Type*} (f : A →* B) (g : B ≃* B') definition pequiv_postcompose {A B B' : Type*} (f : A →* B) (g : B ≃* B')
: pfiber (g ∘* f) ≃* pfiber f := : pfiber (g ∘* f) ≃* pfiber f :=
begin begin
fapply pequiv_of_equiv, esimp, fapply pequiv_of_equiv, esimp,
refine transport_fiber_equiv (g ∘* f) (respect_pt g)⁻¹ ⬝e fiber.equiv_postcompose f g (Point B), refine fiber_equiv_basepoint (g ∘* f) (respect_pt g)⁻¹ ⬝e fiber.equiv_postcompose f g (Point B),
esimp, apply (ap (fiber.mk (Point A))), refine !con.assoc ⬝ _, apply inv_con_eq_of_eq_con, esimp, apply (ap (fiber.mk (Point A))), refine !con.assoc ⬝ _, apply inv_con_eq_of_eq_con,
rewrite [▸*, con.assoc, con.right_inv, con_idp, ap_compose'], rewrite [▸*, con.assoc, con.right_inv, con_idp, ap_compose'],
exact ap_con_eq_con (λ x, ap g⁻¹ᵉ* (ap g (pleft_inv' g x)⁻¹) ⬝ ap g⁻¹ᵉ* (pright_inv g (g x)) ⬝ exact ap_con_eq_con (λ x, ap g⁻¹ᵉ* (ap g (pleft_inv' g x)⁻¹) ⬝ ap g⁻¹ᵉ* (pright_inv g (g x)) ⬝
@ -198,28 +325,6 @@ namespace fiber
{ exact !idp_con⁻¹ } { exact !idp_con⁻¹ }
end end
definition point_fiber_eq {A B : Type} {f : A → B} {b : B} {x y : fiber f b}
(p : point x = point y) (q : point_eq x = ap f p ⬝ point_eq y) :
ap point (fiber_eq p q) = p :=
begin
induction x with a r, induction y with a' s, esimp at *, induction p,
induction q using eq.rec_symm, induction s, reflexivity
end
definition fiber_eq_equiv_fiber {A B : Type} {f : A → B} {b : B} (x y : fiber f b) :
x = y ≃ fiber (ap1_gen f (point_eq x) (point_eq y)) (idpath b) :=
calc
x = y ≃ fiber.sigma_char f b x = fiber.sigma_char f b y :
eq_equiv_fn_eq (fiber.sigma_char f b) x y
... ≃ Σ(p : point x = point y), point_eq x =[p] point_eq y : sigma_eq_equiv
... ≃ Σ(p : point x = point y), (point_eq x)⁻¹ ⬝ ap f p ⬝ point_eq y = idp :
sigma_equiv_sigma_right (λp,
calc point_eq x =[p] point_eq y ≃ point_eq x = ap f p ⬝ point_eq y : eq_pathover_equiv_Fl
... ≃ ap f p ⬝ point_eq y = point_eq x : eq_equiv_eq_symm
... ≃ (point_eq x)⁻¹ ⬝ (ap f p ⬝ point_eq y) = idp : eq_equiv_inv_con_eq_idp
... ≃ (point_eq x)⁻¹ ⬝ ap f p ⬝ point_eq y = idp : equiv_eq_closed_left _ !con.assoc⁻¹)
... ≃ fiber (ap1_gen f (point_eq x) (point_eq y)) (idpath b) : fiber.sigma_char
definition loop_pfiber [constructor] {A B : Type*} (f : A →* B) : Ω (pfiber f) ≃* pfiber (Ω→ f) := definition loop_pfiber [constructor] {A B : Type*} (f : A →* B) : Ω (pfiber f) ≃* pfiber (Ω→ f) :=
pequiv_of_equiv (fiber_eq_equiv_fiber pt pt) pequiv_of_equiv (fiber_eq_equiv_fiber pt pt)
begin begin
@ -229,10 +334,6 @@ namespace fiber
definition pfiber_loop_space {A B : Type*} (f : A →* B) : pfiber (Ω→ f) ≃* Ω (pfiber f) := definition pfiber_loop_space {A B : Type*} (f : A →* B) : pfiber (Ω→ f) ≃* Ω (pfiber f) :=
(loop_pfiber f)⁻¹ᵉ* (loop_pfiber f)⁻¹ᵉ*
definition point_fiber_eq_equiv_fiber {A B : Type} {f : A → B} {b : B} {x y : fiber f b}
(p : x = y) : point (fiber_eq_equiv_fiber x y p) = ap1_gen point idp idp p :=
by induction p; reflexivity
lemma ppoint_loop_pfiber {A B : Type*} (f : A →* B) : lemma ppoint_loop_pfiber {A B : Type*} (f : A →* B) :
ppoint (Ω→ f) ∘* loop_pfiber f ~* Ω→ (ppoint f) := ppoint (Ω→ f) ∘* loop_pfiber f ~* Ω→ (ppoint f) :=
phomotopy.mk (point_fiber_eq_equiv_fiber) phomotopy.mk (point_fiber_eq_equiv_fiber)
@ -288,19 +389,10 @@ namespace fiber
refine !pequiv_postcompose_ppoint⁻¹*, refine !pequiv_postcompose_ppoint⁻¹*,
end end
-- this breaks certain proofs if it is an instance
definition is_trunc_fiber (n : ℕ₋₂) {A B : Type} (f : A → B) (b : B)
[is_trunc n A] [is_trunc (n.+1) B] : is_trunc n (fiber f b) :=
is_trunc_equiv_closed_rev n !fiber.sigma_char _
definition is_trunc_pfiber (n : ℕ₋₂) {A B : Type*} (f : A →* B) definition is_trunc_pfiber (n : ℕ₋₂) {A B : Type*} (f : A →* B)
[is_trunc n A] [is_trunc (n.+1) B] : is_trunc n (pfiber f) := [is_trunc n A] [is_trunc (n.+1) B] : is_trunc n (pfiber f) :=
is_trunc_fiber n f pt is_trunc_fiber n f pt
definition fiber_equiv_of_is_contr [constructor] {A B : Type} (f : A → B) (b : B) [is_contr B] :
fiber f b ≃ A :=
!fiber.sigma_char ⬝e !sigma_equiv_of_is_contr_right
definition pfiber_pequiv_of_is_contr [constructor] {A B : Type*} (f : A →* B) [is_contr B] : definition pfiber_pequiv_of_is_contr [constructor] {A B : Type*} (f : A →* B) [is_contr B] :
pfiber f ≃* A := pfiber f ≃* A :=
pequiv_of_equiv (fiber_equiv_of_is_contr f pt) idp pequiv_of_equiv (fiber_equiv_of_is_contr f pt) idp
@ -316,7 +408,7 @@ namespace fiber
definition pfiber_ppoint_pequiv {A B : Type*} (f : A →* B) : pfiber (ppoint f) ≃* Ω B := definition pfiber_ppoint_pequiv {A B : Type*} (f : A →* B) : pfiber (ppoint f) ≃* Ω B :=
pequiv_of_equiv (pfiber_ppoint_equiv f) !con.left_inv pequiv_of_equiv (pfiber_ppoint_equiv f) !con.left_inv
definition fiber_ppoint_equiv_eq {A B : Type*} {f : A →* B} {a : A} (p : f a = pt) definition pfiber_ppoint_equiv_eq {A B : Type*} {f : A →* B} {a : A} (p : f a = pt)
(q : ppoint f (fiber.mk a p) = pt) : (q : ppoint f (fiber.mk a p) = pt) :
pfiber_ppoint_equiv f (fiber.mk (fiber.mk a p) q) = (respect_pt f)⁻¹ ⬝ ap f q⁻¹ ⬝ p := pfiber_ppoint_equiv f (fiber.mk (fiber.mk a p) q) = (respect_pt f)⁻¹ ⬝ ap f q⁻¹ ⬝ p :=
begin begin
@ -328,58 +420,53 @@ namespace fiber
refine ap_compose f pr₁ _ ⬝ ap02 f !ap_pr1_center_eq_sigma_eq', refine ap_compose f pr₁ _ ⬝ ap02 f !ap_pr1_center_eq_sigma_eq',
end end
definition fiber_ppoint_equiv_inv_eq {A B : Type*} (f : A →* B) (p : Ω B) : definition pfiber_ppoint_equiv_inv_eq {A B : Type*} (f : A →* B) (p : Ω B) :
(pfiber_ppoint_equiv f)⁻¹ᵉ p = fiber.mk (fiber.mk pt (respect_pt f ⬝ p)) idp := (pfiber_ppoint_equiv f)⁻¹ᵉ p = fiber.mk (fiber.mk pt (respect_pt f ⬝ p)) idp :=
begin begin
apply inv_eq_of_eq, apply inv_eq_of_eq,
refine _ ⬝ !fiber_ppoint_equiv_eq⁻¹, refine _ ⬝ !pfiber_ppoint_equiv_eq⁻¹,
exact !inv_con_cancel_left⁻¹ exact !inv_con_cancel_left⁻¹
end end
definition loopn_pfiber [constructor] {A B : Type*} (n : ) (f : A →* B) :
Ω[n] (pfiber f) ≃* pfiber (Ω→[n] f) :=
begin
induction n with n IH, reflexivity, exact loop_pequiv_loop IH ⬝e* loop_pfiber (Ω→[n] f),
end
definition is_contr_pfiber_pid (A : Type*) : is_contr (pfiber (pid A)) :=
is_contr_fiber_id A pt
definition pfiber_functor [constructor] {A A' B B' : Type*} {f : A →* B} {f' : A' →* B'}
(g : A →* A') (h : B →* B') (H : psquare g h f f') : pfiber f →* pfiber f' :=
pmap.mk (fiber_functor g h H (respect_pt h))
begin
fapply fiber_eq,
exact respect_pt g,
exact !con.assoc ⬝ to_homotopy_pt H
end
definition ppoint_natural {A A' B B' : Type*} {f : A →* B} {f' : A' →* B'}
(g : A →* A') (h : B →* B') (H : psquare g h f f') :
psquare (ppoint f) (ppoint f') (pfiber_functor g h H) g :=
begin
fapply phomotopy.mk,
{ intro x, reflexivity },
{ refine !idp_con ⬝ _ ⬝ !idp_con⁻¹, esimp, apply point_fiber_eq }
end
/- A less commonly used pointed fiber with basepoint (f a) for some a in the domain of f -/
definition pointed_fiber [constructor] (f : A → B) (a : A) : Type* :=
pointed.Mk (fiber.mk a (idpath (f a)))
end fiber end fiber
open fiber
open function is_equiv /- A function is truncated if it has truncated fibers -/
definition is_trunc_fun [reducible] (n : ℕ₋₂) (f : A → B) :=
Π(b : B), is_trunc n (fiber f b)
namespace fiber definition is_contr_fun [reducible] (f : A → B) := is_trunc_fun -2 f
/- Theorem 4.7.6 -/
variables {A : Type} {P Q : A → Type}
variable (f : Πa, P a → Q a)
definition fiber_total_equiv [constructor] {a : A} (q : Q a) definition is_trunc_fun_id (k : ℕ₋₂) (A : Type) : is_trunc_fun k (@id A) :=
: fiber (total f) ⟨a , q⟩ ≃ fiber (f a) q := λa, is_trunc_of_is_contr _ _ !is_contr_fiber_id
calc
fiber (total f) ⟨a , q⟩
≃ Σ(w : Σx, P x), ⟨w.1 , f w.1 w.2 ⟩ = ⟨a , q⟩
: fiber.sigma_char
... ≃ Σ(x : A), Σ(p : P x), ⟨x , f x p⟩ = ⟨a , q⟩
: sigma_assoc_equiv
... ≃ Σ(x : A), Σ(p : P x), Σ(H : x = a), f x p =[H] q
:
begin
apply sigma_equiv_sigma_right, intro x,
apply sigma_equiv_sigma_right, intro p,
apply sigma_eq_equiv
end
... ≃ Σ(x : A), Σ(H : x = a), Σ(p : P x), f x p =[H] q
:
begin
apply sigma_equiv_sigma_right, intro x,
apply sigma_comm_equiv
end
... ≃ Σ(w : Σx, x = a), Σ(p : P w.1), f w.1 p =[w.2] q
: sigma_assoc_equiv
... ≃ Σ(p : P (center (Σx, x=a)).1), f (center (Σx, x=a)).1 p =[(center (Σx, x=a)).2] q
: sigma_equiv_of_is_contr_left
... ≃ Σ(p : P a), f a p =[idpath a] q
: equiv_of_eq idp
... ≃ Σ(p : P a), f a p = q
:
begin
apply sigma_equiv_sigma_right, intro p,
apply pathover_idp
end
... ≃ fiber (f a) q
: fiber.sigma_char
end fiber

View file

@ -532,7 +532,7 @@ begin
end end
definition fin_two_equiv_bool : fin 2 ≃ bool := definition fin_two_equiv_bool : fin 2 ≃ bool :=
let H := equiv_unit_of_is_contr (fin 1) in let H := equiv_unit_of_is_contr (fin 1) _ in
calc calc
fin 2 ≃ fin (1 + 1) : equiv.refl fin 2 ≃ fin (1 + 1) : equiv.refl
... ≃ fin 1 + fin 1 : fin_sum_equiv ... ≃ fin 1 + fin 1 : fin_sum_equiv
@ -540,7 +540,7 @@ calc
... ≃ bool : bool_equiv_unit_sum_unit ... ≃ bool : bool_equiv_unit_sum_unit
definition fin_sum_unit_equiv (n : nat) : fin n + unit ≃ fin (nat.succ n) := definition fin_sum_unit_equiv (n : nat) : fin n + unit ≃ fin (nat.succ n) :=
let H := equiv_unit_of_is_contr (fin 1) in let H := equiv_unit_of_is_contr (fin 1) _ in
calc calc
fin n + unit ≃ fin n + fin 1 : H fin n + unit ≃ fin n + fin 1 : H
... ≃ fin (nat.succ n) : fin_sum_equiv ... ≃ fin (nat.succ n) : fin_sum_equiv

View file

@ -36,6 +36,9 @@ namespace int
end end
definition not_neg_succ_le_of_nat {n m : } : ¬m ≤ -[1+n] :=
by cases m: exact id
definition is_equiv_succ [constructor] [instance] : is_equiv succ := definition is_equiv_succ [constructor] [instance] : is_equiv succ :=
adjointify succ pred (λa, !add_sub_cancel) (λa, !sub_add_cancel) adjointify succ pred (λa, !add_sub_cancel) (λa, !sub_add_cancel)
definition equiv_succ [constructor] : := equiv.mk succ _ definition equiv_succ [constructor] : := equiv.mk succ _
@ -60,6 +63,15 @@ namespace int
lemma le_of_max0_le {n : } {m : } (h : max0 n ≤ m) : n ≤ of_nat m := lemma le_of_max0_le {n : } {m : } (h : max0 n ≤ m) : n ≤ of_nat m :=
le.trans (le_max0 n) (of_nat_le_of_nat_of_le h) le.trans (le_max0 n) (of_nat_le_of_nat_of_le h)
definition max0_monotone {n m : } (H : n ≤ m) : max0 n ≤ max0 m :=
begin
induction n with n n,
{ induction m with m m,
{ exact le_of_of_nat_le_of_nat H },
{ exfalso, exact not_neg_succ_le_of_nat H }},
{ apply zero_le }
end
-- definition iterate_trans {A : Type} (f : A ≃ A) (a : ) -- definition iterate_trans {A : Type} (f : A ≃ A) (a : )
-- : iterate f a ⬝e f = iterate f (a + 1) := -- : iterate f a ⬝e f = iterate f (a + 1) :=
-- sorry -- sorry

View file

@ -129,9 +129,7 @@ namespace lift
definition is_embedding_lift [instance] : is_embedding lift := definition is_embedding_lift [instance] : is_embedding lift :=
begin begin
intro A A', fapply is_equiv.homotopy_closed, intro A A', refine is_equiv_of_equiv_of_homotopy !lift_eq_lift_equiv⁻¹ᵉ _,
exact to_inv !lift_eq_lift_equiv,
exact _,
{ intro p, induction p, { intro p, induction p,
esimp [lift_eq_lift_equiv,equiv.trans,equiv.symm,eq_equiv_equiv], esimp [lift_eq_lift_equiv,equiv.trans,equiv.symm,eq_equiv_equiv],
rewrite [equiv_of_eq_refl, lift_equiv_lift_refl], rewrite [equiv_of_eq_refl, lift_equiv_lift_refl],

View file

@ -6,7 +6,7 @@ Author: Floris van Doorn
Theorems about the natural numbers specific to HoTT Theorems about the natural numbers specific to HoTT
-/ -/
import .order types.pointed .sub import .sub
open is_trunc unit empty eq equiv algebra pointed is_equiv equiv function open is_trunc unit empty eq equiv algebra pointed is_equiv equiv function
@ -30,9 +30,9 @@ namespace nat
definition is_prop_lt [instance] (n m : ) : is_prop (n < m) := !is_prop_le definition is_prop_lt [instance] (n m : ) : is_prop (n < m) := !is_prop_le
definition le_equiv_succ_le_succ (n m : ) : (n ≤ m) ≃ (succ n ≤ succ m) := definition le_equiv_succ_le_succ (n m : ) : (n ≤ m) ≃ (succ n ≤ succ m) :=
equiv_of_is_prop succ_le_succ le_of_succ_le_succ equiv_of_is_prop succ_le_succ le_of_succ_le_succ _ _
definition le_succ_equiv_pred_le (n m : ) : (n ≤ succ m) ≃ (pred n ≤ m) := definition le_succ_equiv_pred_le (n m : ) : (n ≤ succ m) ≃ (pred n ≤ m) :=
equiv_of_is_prop pred_le_pred le_succ_of_pred_le equiv_of_is_prop pred_le_pred le_succ_of_pred_le _ _
theorem lt_by_cases_lt {a b : } {P : Type} (H1 : a < b → P) (H2 : a = b → P) theorem lt_by_cases_lt {a b : } {P : Type} (H1 : a < b → P) (H2 : a = b → P)
(H3 : a > b → P) (H : a < b) : lt.by_cases H1 H2 H3 = H1 H := (H3 : a > b → P) (H : a < b) : lt.by_cases H1 H2 H3 = H1 H :=
@ -267,7 +267,7 @@ namespace nat
definition iterate_equiv {A : Type} (f : A ≃ A) (n : ) : A ≃ A := definition iterate_equiv {A : Type} (f : A ≃ A) (n : ) : A ≃ A :=
equiv.mk (iterate f n) equiv.mk (iterate f n)
(by induction n with n IH; apply is_equiv_id; exact is_equiv_compose f (iterate f n)) (by induction n with n IH; apply is_equiv_id; exact is_equiv_compose f (iterate f n) _ _)
definition iterate_equiv2 {A : Type} {C : A → Type} (f : A → A) (h : Πa, C a ≃ C (f a)) definition iterate_equiv2 {A : Type} {C : A → Type} (f : A → A) (h : Πa, C a ≃ C (f a))
(k : ) (a : A) : C a ≃ C (f^[k] a) := (k : ) (a : A) : C a ≃ C (f^[k] a) :=

View file

@ -9,7 +9,7 @@ The basic definitions are in init.pointed
See also .pointed2 See also .pointed2
-/ -/
import .nat.basic ..arity ..prop_trunc import .nat.basic ..prop_trunc
open is_trunc eq prod sigma nat equiv option is_equiv bool unit sigma.ops sum algebra function open is_trunc eq prod sigma nat equiv option is_equiv bool unit sigma.ops sum algebra function
namespace pointed namespace pointed
@ -201,6 +201,8 @@ namespace pointed
definition ppmap [constructor] (A B : Type*) : Type* := definition ppmap [constructor] (A B : Type*) : Type* :=
@pppi A (λa, B) @pppi A (λa, B)
infixr ` →** `:29 := ppmap
definition pcast [constructor] {A B : Type*} (p : A = B) : A →* B := definition pcast [constructor] {A B : Type*} (p : A = B) : A →* B :=
pmap.mk (cast (ap pType.carrier p)) (by induction p; reflexivity) pmap.mk (cast (ap pType.carrier p)) (by induction p; reflexivity)
@ -297,7 +299,7 @@ namespace pointed
definition is_equiv_ap1 (f : A →* B) [is_equiv f] : is_equiv (ap1 f) := definition is_equiv_ap1 (f : A →* B) [is_equiv f] : is_equiv (ap1 f) :=
begin begin
induction B with B b, induction f with f pf, esimp at *, cases pf, esimp, induction B with B b, induction f with f pf, esimp at *, cases pf, esimp,
apply is_equiv.homotopy_closed (ap f), refine is_equiv.homotopy_closed (ap f) _ _,
intro p, exact !idp_con⁻¹ intro p, exact !idp_con⁻¹
end end
@ -830,7 +832,7 @@ namespace pointed
pequiv_of_pmap (ptransport B p) !is_equiv_tr pequiv_of_pmap (ptransport B p) !is_equiv_tr
definition pequiv_change_fun [constructor] (f : A ≃* B) (f' : A →* B) (Heq : f ~ f') : A ≃* B := definition pequiv_change_fun [constructor] (f : A ≃* B) (f' : A →* B) (Heq : f ~ f') : A ≃* B :=
pequiv_of_pmap f' (is_equiv.homotopy_closed f Heq) pequiv_of_pmap f' (is_equiv.homotopy_closed f Heq _)
definition pequiv_change_inv [constructor] (f : A ≃* B) (f' : B →* A) (Heq : to_pinv f ~ f') definition pequiv_change_inv [constructor] (f : A ≃* B) (f' : B →* A) (Heq : to_pinv f ~ f')
: A ≃* B := : A ≃* B :=
@ -1155,7 +1157,7 @@ namespace pointed
Ω[succ n](pointed.Mk p) = Ω[n](Ω (pointed.Mk p)) : eq_of_pequiv !loopn_succ_in Ω[succ n](pointed.Mk p) = Ω[n](Ω (pointed.Mk p)) : eq_of_pequiv !loopn_succ_in
... = Ω[n] (Ω[2] A) : loopn_loop_irrel ... = Ω[n] (Ω[2] A) : loopn_loop_irrel
... = Ω[2+n] A : eq_of_pequiv !loopn_add ... = Ω[2+n] A : eq_of_pequiv !loopn_add
... = Ω[n+2] A : by rewrite [algebra.add.comm] ... = Ω[n+2] A : by rewrite [nat.add_comm]
section psquare section psquare
/- /-

View file

@ -110,6 +110,16 @@ namespace sigma
: (u = v) ≃ (Σ(p : u.1 = v.1), u.2 =[p] v.2) := : (u = v) ≃ (Σ(p : u.1 = v.1), u.2 =[p] v.2) :=
(equiv.mk sigma_eq_unc _)⁻¹ᵉ (equiv.mk sigma_eq_unc _)⁻¹ᵉ
/- induction principle for a path between pairs -/
definition eq.rec_sigma {A : Type} {B : A → Type} {a : A} {b : B a}
(P : Π⦃a'⦄ {b' : B a'}, ⟨a, b⟩ = ⟨a', b'⟩ → Type)
(IH : P idp) ⦃a' : A⦄ {b' : B a'} (p : ⟨a, b⟩ = ⟨a', b'⟩) : P p :=
begin
apply transport (λp, P p) (to_left_inv !sigma_eq_equiv p),
generalize !sigma_eq_equiv p, esimp, intro q,
induction q with q₁ q₂, induction q₂, exact IH
end
definition dpair_eq_dpair_con (p1 : a = a' ) (q1 : b =[p1] b' ) definition dpair_eq_dpair_con (p1 : a = a' ) (q1 : b =[p1] b' )
(p2 : a' = a'') (q2 : b' =[p2] b'') : (p2 : a' = a'') (q2 : b' =[p2] b'') :
dpair_eq_dpair (p1 ⬝ p2) (q1 ⬝o q2) = dpair_eq_dpair p1 q1 ⬝ dpair_eq_dpair p2 q2 := dpair_eq_dpair (p1 ⬝ p2) (q1 ⬝o q2) = dpair_eq_dpair p1 q1 ⬝ dpair_eq_dpair p2 q2 :=
@ -120,8 +130,13 @@ namespace sigma
sigma_eq (p1 ⬝ p2) (q1 ⬝o q2) = sigma_eq p1 q1 ⬝ sigma_eq p2 q2 := sigma_eq (p1 ⬝ p2) (q1 ⬝o q2) = sigma_eq p1 q1 ⬝ sigma_eq p2 q2 :=
by induction u; induction v; induction w; apply dpair_eq_dpair_con by induction u; induction v; induction w; apply dpair_eq_dpair_con
definition dpair_eq_dpair_inv {A : Type} {B : A → Type} {a a' : A} {b : B a} {b' : B a'} (p : a = a') definition sigma_eq_concato_eq {b : B a} {b₁ b₂ : B a'}
(q : b =[p] b') : (dpair_eq_dpair p q)⁻¹ = dpair_eq_dpair p⁻¹ q⁻¹ᵒ := (p : a = a') (q : b =[p] b₁) (q' : b₁ = b₂) :
sigma_eq p (q ⬝op q') = sigma_eq p q ⬝ ap (dpair a') q' :=
by induction q'; reflexivity
definition dpair_eq_dpair_inv (p : a = a') (q : b =[p] b') :
(dpair_eq_dpair p q)⁻¹ = dpair_eq_dpair p⁻¹ q⁻¹ᵒ :=
begin induction q, reflexivity end begin induction q, reflexivity end
local attribute dpair_eq_dpair [reducible] local attribute dpair_eq_dpair [reducible]
@ -130,6 +145,23 @@ namespace sigma
dpair_eq_dpair idp (pathover_idp_of_eq (tr_eq_of_pathover q)) := dpair_eq_dpair idp (pathover_idp_of_eq (tr_eq_of_pathover q)) :=
by induction q; reflexivity by induction q; reflexivity
definition ap_sigma_pr1 {A B : Type} {C : B → Type} {a₁ a₂ : A} (f : A → B) (g : Πa, C (f a))
(p : a₁ = a₂) : (ap (λa, ⟨f a, g a⟩) p)..1 = ap f p :=
by induction p; reflexivity
definition ap_sigma_pr2 {A B : Type} {C : B → Type} {a₁ a₂ : A} (f : A → B) (g : Πa, C (f a))
(p : a₁ = a₂) : (ap (λa, ⟨f a, g a⟩) p)..2 =
change_path (ap_sigma_pr1 f g p)⁻¹ (pathover_ap C f (apd g p)) :=
by induction p; reflexivity
definition sigma_eq_pr2_constant {A B : Type} {a a' : A} {b b' : B} (p : a = a')
(q : b =[p] b') : ap pr2 (sigma_eq p q) = (eq_of_pathover q) :=
by induction q; reflexivity
definition sigma_eq_pr2_constant2 {A B : Type} {a a' : A} {b b' : B} (p : a = a')
(q : b = b') : ap pr2 (sigma_eq p (pathover_of_eq p q)) = q :=
by induction p; induction q; reflexivity
/- eq_pr1 commutes with the groupoid structure. -/ /- eq_pr1 commutes with the groupoid structure. -/
definition eq_pr1_idp (u : Σa, B a) : (refl u) ..1 = refl (u.1) := idp definition eq_pr1_idp (u : Σa, B a) : (refl u) ..1 = refl (u.1) := idp
@ -165,10 +197,16 @@ namespace sigma
definition sigma_eq2_unc {p q : u = v} (rs : Σ(r : p..1 = q..1), p..2 =[r] q..2) : p = q := definition sigma_eq2_unc {p q : u = v} (rs : Σ(r : p..1 = q..1), p..2 =[r] q..2) : p = q :=
destruct rs sigma_eq2 destruct rs sigma_eq2
/- two equalities about applying a function to a pair of equalities. These two functions are not
definitionally equal (but they are equal on all pairs) -/
definition ap_dpair_eq_dpair (f : Πa, B a → A') (p : a = a') (q : b =[p] b') definition ap_dpair_eq_dpair (f : Πa, B a → A') (p : a = a') (q : b =[p] b')
: ap (sigma.rec f) (dpair_eq_dpair p q) = apd011 f p q := : ap (sigma.rec f) (dpair_eq_dpair p q) = apd011 f p q :=
by induction q; reflexivity by induction q; reflexivity
definition ap_dpair_eq_dpair_pr (f : Πa, B a → A') (p : a = a') (q : b =[p] b') :
ap (λx, f x.1 x.2) (dpair_eq_dpair p q) = apd011 f p q :=
by induction q; reflexivity
/- Transport -/ /- Transport -/
/- The concrete description of transport in sigmas (and also pis) is rather trickier than in the other types. In particular, these cannot be described just in terms of transport in simpler types; they require also the dependent transport [transportD]. /- The concrete description of transport in sigmas (and also pis) is rather trickier than in the other types. In particular, these cannot be described just in terms of transport in simpler types; they require also the dependent transport [transportD].
@ -255,6 +293,41 @@ namespace sigma
definition total [reducible] [unfold 5] {B' : A → Type} (g : Πa, B a → B' a) : (Σa, B a) → (Σa, B' a) := definition total [reducible] [unfold 5] {B' : A → Type} (g : Πa, B a → B' a) : (Σa, B a) → (Σa, B' a) :=
sigma_functor id g sigma_functor id g
definition sigma_functor_compose {A A' A'' : Type} {B : A → Type} {B' : A' → Type}
{B'' : A'' → Type} {f' : A' → A''} {f : A → A'} (g' : Πa, B' a → B'' (f' a))
(g : Πa, B a → B' (f a)) (x : Σa, B a) :
sigma_functor f' g' (sigma_functor f g x) = sigma_functor (f' ∘ f) (λa, g' (f a) ∘ g a) x :=
begin
reflexivity
end
definition sigma_functor_homotopy {A A' : Type} {B : A → Type} {B' : A' → Type}
{f f' : A → A'} {g : Πa, B a → B' (f a)} {g' : Πa, B a → B' (f' a)} (h : f ~ f')
(k : Πa b, g a b =[h a] g' a b) (x : Σa, B a) : sigma_functor f g x = sigma_functor f' g' x :=
sigma_eq (h x.1) (k x.1 x.2)
section hsquare
variables {A₀₀ A₂₀ A₀₂ A₂₂ : Type}
{B₀₀ : A₀₀ → Type} {B₂₀ : A₂₀ → Type} {B₀₂ : A₀₂ → Type} {B₂₂ : A₂₂ → Type}
{f₁₀ : A₀₀ → A₂₀} {f₁₂ : A₀₂ → A₂₂} {f₀₁ : A₀₀ → A₀₂} {f₂₁ : A₂₀ → A₂₂}
{g₁₀ : Πa, B₀₀ a → B₂₀ (f₁₀ a)} {g₁₂ : Πa, B₀₂ a → B₂₂ (f₁₂ a)}
{g₀₁ : Πa, B₀₀ a → B₀₂ (f₀₁ a)} {g₂₁ : Πa, B₂₀ a → B₂₂ (f₂₁ a)}
definition sigma_functor_hsquare (h : hsquare f₁₀ f₁₂ f₀₁ f₂₁)
(k : Πa (b : B₀₀ a), g₂₁ _ (g₁₀ _ b) =[h a] g₁₂ _ (g₀₁ _ b)) :
hsquare (sigma_functor f₁₀ g₁₀) (sigma_functor f₁₂ g₁₂)
(sigma_functor f₀₁ g₀₁) (sigma_functor f₂₁ g₂₁) :=
λx, sigma_functor_compose g₂₁ g₁₀ x ⬝
sigma_functor_homotopy h k x ⬝
(sigma_functor_compose g₁₂ g₀₁ x)⁻¹
end hsquare
definition sigma_functor2 [constructor] {A₁ A₂ A₃ : Type}
{B₁ : A₁ → Type} {B₂ : A₂ → Type} {B₃ : A₃ → Type}
(f : A₁ → A₂ → A₃) (g : Π⦃a₁ a₂⦄, B₁ a₁ → B₂ a₂ → B₃ (f a₁ a₂))
(x₁ : Σa₁, B₁ a₁) (x₂ : Σa₂, B₂ a₂) : Σa₃, B₃ a₃ :=
⟨f x₁.1 x₂.1, g x₁.2 x₂.2⟩
/- Equivalences -/ /- Equivalences -/
definition is_equiv_sigma_functor [constructor] [H1 : is_equiv f] [H2 : Π a, is_equiv (g a)] definition is_equiv_sigma_functor [constructor] [H1 : is_equiv f] [H2 : Π a, is_equiv (g a)]
: is_equiv (sigma_functor f g) := : is_equiv (sigma_functor f g) :=
@ -295,8 +368,14 @@ namespace sigma
definition sigma_equiv_sigma_left' [constructor] (Hf : A' ≃ A) : (Σa, B (Hf a)) ≃ (Σa', B a') := definition sigma_equiv_sigma_left' [constructor] (Hf : A' ≃ A) : (Σa, B (Hf a)) ≃ (Σa', B a') :=
sigma_equiv_sigma Hf (λa, erfl) sigma_equiv_sigma Hf (λa, erfl)
definition ap_sigma_functor_eq_dpair (p : a = a') (q : b =[p] b') : definition ap_sigma_functor_sigma_eq {A A' : Type} {B : A → Type} {B' : A' → Type} {a a' : A}
ap (sigma_functor f g) (sigma_eq p q) = sigma_eq (ap f p) (pathover.rec_on q idpo) := {b : B a} {b' : B a'} (f : A → A') (g : Πa, B a → B' (f a)) (p : a = a') (q : b =[p] b') :
ap (sigma_functor f g) (sigma_eq p q) = sigma_eq (ap f p) (pathover_ap B' f (apo g q)) :=
by induction q; reflexivity
definition ap_sigma_functor_id_sigma_eq {A : Type} {B B' : A → Type}
{a a' : A} {b : B a} {b' : B a'} (g : Πa, B a → B' a) (p : a = a') (q : b =[p] b') :
ap (sigma_functor id g) (sigma_eq p q) = sigma_eq p (apo g q) :=
by induction q; reflexivity by induction q; reflexivity
definition sigma_ua {A B : Type} (C : A ≃ B → Type) : definition sigma_ua {A B : Type} (C : A ≃ B → Type) :
@ -499,7 +578,7 @@ namespace sigma
definition is_equiv_subtype_eq [constructor] [H : Πa, is_prop (B a)] (u v : {a | B a}) definition is_equiv_subtype_eq [constructor] [H : Πa, is_prop (B a)] (u v : {a | B a})
: is_equiv (subtype_eq : u.1 = v.1 → u = v) := : is_equiv (subtype_eq : u.1 = v.1 → u = v) :=
!is_equiv_compose is_equiv_compose _ _ _ _
local attribute is_equiv_subtype_eq [instance] local attribute is_equiv_subtype_eq [instance]
definition equiv_subtype [constructor] [H : Πa, is_prop (B a)] (u v : {a | B a}) : definition equiv_subtype [constructor] [H : Πa, is_prop (B a)] (u v : {a | B a}) :
@ -520,7 +599,7 @@ namespace sigma
_ _
/- truncatedness -/ /- truncatedness -/
theorem is_trunc_sigma (B : A → Type) (n : trunc_index) theorem is_trunc_sigma (B : A → Type) (n : ℕ₋₂)
[HA : is_trunc n A] [HB : Πa, is_trunc n (B a)] : is_trunc n (Σa, B a) := [HA : is_trunc n A] [HB : Πa, is_trunc n (B a)] : is_trunc n (Σa, B a) :=
begin begin
revert A B HA HB, revert A B HA HB,
@ -530,9 +609,9 @@ namespace sigma
exact is_trunc_equiv_closed_rev n !sigma_eq_equiv (IH _ _ _ _) } exact is_trunc_equiv_closed_rev n !sigma_eq_equiv (IH _ _ _ _) }
end end
theorem is_trunc_subtype (B : A → Prop) (n : trunc_index) theorem is_trunc_subtype (B : A → Prop) (n : ℕ₋₂)
[HA : is_trunc (n.+1) A] : is_trunc (n.+1) (Σa, B a) := [HA : is_trunc (n.+1) A] : is_trunc (n.+1) (Σa, B a) :=
@(is_trunc_sigma B (n.+1)) _ (λa, !is_trunc_succ_of_is_prop) @(is_trunc_sigma B (n.+1)) _ (λa, is_trunc_succ_of_is_prop _ _ _)
/- if the total space is a mere proposition, you can equate two points in the base type by /- if the total space is a mere proposition, you can equate two points in the base type by
finding points in their fibers -/ finding points in their fibers -/

View file

@ -285,7 +285,7 @@ namespace is_trunc
induction n with n, induction n with n,
{exfalso, exact not_succ_le_minus_two Hn}, {exfalso, exact not_succ_le_minus_two Hn},
{apply is_trunc_succ_intro, intro a a', {apply is_trunc_succ_intro, intro a a',
fapply @is_trunc_is_equiv_closed_rev _ _ n (ap f)} exact is_trunc_is_equiv_closed_rev n (ap f) _ _ }
end end
theorem is_trunc_is_retraction_closed (f : A → B) [Hf : is_retraction f] theorem is_trunc_is_retraction_closed (f : A → B) [Hf : is_retraction f]
@ -336,9 +336,9 @@ namespace is_trunc
apply is_trunc_equiv_closed_rev _ !trunctype_eq_equiv, apply is_trunc_equiv_closed_rev _ !trunctype_eq_equiv,
apply is_trunc_equiv_closed_rev _ !eq_equiv_equiv, apply is_trunc_equiv_closed_rev _ !eq_equiv_equiv,
induction n, induction n,
{ apply @is_contr_of_inhabited_prop, { apply is_contr_of_inhabited_prop,
{ apply is_trunc_equiv }, { exact equiv_of_is_contr_of_is_contr _ _ },
{ apply equiv_of_is_contr_of_is_contr}}, { apply is_trunc_equiv }},
{ apply is_trunc_equiv } { apply is_trunc_equiv }
end end
@ -411,7 +411,7 @@ namespace is_trunc
(mere : Π(a b : A), is_prop (R a b)) (refl : Π(a : A), R a a) (mere : Π(a b : A), is_prop (R a b)) (refl : Π(a : A), R a a)
(imp : Π{a b : A}, R a b → a = b) (a b : A) : R a b ≃ a = b := (imp : Π{a b : A}, R a b → a = b) (a b : A) : R a b ≃ a = b :=
have is_set A, from is_set_of_relation R mere refl @imp, have is_set A, from is_set_of_relation R mere refl @imp,
equiv_of_is_prop imp (λp, p ▸ refl a) equiv_of_is_prop imp (λp, p ▸ refl a) _ _
local attribute not [reducible] local attribute not [reducible]
theorem is_set_of_double_neg_elim {A : Type} (H : Π(a b : A), ¬¬a = b → a = b) theorem is_set_of_double_neg_elim {A : Type} (H : Π(a b : A), ¬¬a = b → a = b)
@ -468,7 +468,7 @@ namespace is_trunc
begin begin
induction n with n, induction n with n,
{ esimp [sub_two,loopn], apply iff.intro, { esimp [sub_two,loopn], apply iff.intro,
intro H a, exact is_contr_of_inhabited_prop a, intro H a, exact is_contr_of_inhabited_prop a _,
intro H, apply is_prop_of_imp_is_contr, exact H}, intro H, apply is_prop_of_imp_is_contr, exact H},
{ apply is_trunc_iff_is_contr_loop_succ}, { apply is_trunc_iff_is_contr_loop_succ},
end end
@ -512,7 +512,7 @@ namespace is_trunc
transport (λk, is_trunc k A) p H transport (λk, is_trunc k A) p H
definition pequiv_punit_of_is_contr [constructor] (A : Type*) (H : is_contr A) : A ≃* punit := definition pequiv_punit_of_is_contr [constructor] (A : Type*) (H : is_contr A) : A ≃* punit :=
pequiv_of_equiv (equiv_unit_of_is_contr A) (@is_prop.elim unit _ _ _) pequiv_of_equiv (equiv_unit_of_is_contr A _) (@is_prop.elim unit _ _ _)
definition pequiv_punit_of_is_contr' [constructor] (A : Type) (H : is_contr A) definition pequiv_punit_of_is_contr' [constructor] (A : Type) (H : is_contr A)
: pointed.MK A (center A) ≃* punit := : pointed.MK A (center A) ≃* punit :=
@ -522,9 +522,9 @@ namespace is_trunc
(b : B) [is_trunc n A] [is_trunc n B] : is_trunc n (is_contr (fiber f b)) := (b : B) [is_trunc n A] [is_trunc n B] : is_trunc n (is_contr (fiber f b)) :=
begin begin
cases n, cases n,
{ apply is_contr_of_inhabited_prop, apply is_contr_fun_of_is_equiv, { refine is_contr_of_inhabited_prop _ _, apply is_contr_fun_of_is_equiv,
apply is_equiv_of_is_contr }, exact is_equiv_of_is_contr _ _ _ },
{ apply is_trunc_succ_of_is_prop } { exact is_trunc_succ_of_is_prop _ _ _ }
end end
end is_trunc open is_trunc end is_trunc open is_trunc
@ -650,12 +650,12 @@ namespace trunc
begin begin
revert A m H, eapply (trunc_index.rec_on n), revert A m H, eapply (trunc_index.rec_on n),
{ clear n, intro A m H, refine is_contr_equiv_closed_rev _ H, { clear n, intro A m H, refine is_contr_equiv_closed_rev _ H,
{ apply trunc_equiv, apply (@is_trunc_of_le _ -2), apply minus_two_le} }, { apply trunc_equiv, exact is_trunc_of_is_contr _ _ _ } },
{ clear n, intro n IH A m H, induction m with m, { clear n, intro n IH A m H, induction m with m,
{ apply (@is_trunc_of_le _ -2), apply minus_two_le}, { exact is_trunc_of_is_contr _ _ _ },
{ apply is_trunc_succ_intro, intro aa aa', { apply is_trunc_succ_intro, intro aa aa',
apply (@trunc.rec_on _ _ _ aa (λy, !is_trunc_succ_of_is_prop)), apply (@trunc.rec_on _ _ _ aa (λy, is_trunc_succ_of_is_prop _ _ _)),
eapply (@trunc.rec_on _ _ _ aa' (λy, !is_trunc_succ_of_is_prop)), eapply (@trunc.rec_on _ _ _ aa' (λy, is_trunc_succ_of_is_prop _ _ _)),
intro a a', apply is_trunc_equiv_closed_rev _ !tr_eq_tr_equiv (IH _ _ _) }} intro a a', apply is_trunc_equiv_closed_rev _ !tr_eq_tr_equiv (IH _ _ _) }}
end end
@ -663,7 +663,7 @@ namespace trunc
definition trunc_trunc_equiv_left [constructor] (A : Type) {n m : ℕ₋₂} (H : n ≤ m) definition trunc_trunc_equiv_left [constructor] (A : Type) {n m : ℕ₋₂} (H : n ≤ m)
: trunc n (trunc m A) ≃ trunc n A := : trunc n (trunc m A) ≃ trunc n A :=
begin begin
note H2 := is_trunc_of_le (trunc n A) H, note H2 := is_trunc_of_le (trunc n A) H _,
fapply equiv.MK, fapply equiv.MK,
{ intro x, induction x with x, induction x with x, exact tr x }, { intro x, induction x with x, induction x with x, exact tr x },
{ exact trunc_functor n tr }, { exact trunc_functor n tr },
@ -675,7 +675,7 @@ namespace trunc
: trunc m (trunc n A) ≃ trunc n A := : trunc m (trunc n A) ≃ trunc n A :=
begin begin
apply trunc_equiv, apply trunc_equiv,
exact is_trunc_of_le _ H, exact is_trunc_of_le _ H _,
end end
definition trunc_equiv_trunc_of_le {n m : ℕ₋₂} {A B : Type} (H : n ≤ m) definition trunc_equiv_trunc_of_le {n m : ℕ₋₂} {A B : Type} (H : n ≤ m)
@ -734,7 +734,7 @@ namespace trunc
begin begin
cases n with n: intro b, cases n with n: intro b,
{ exact tr (fiber.mk !center !is_prop.elim)}, { exact tr (fiber.mk !center !is_prop.elim)},
{ refine @trunc.rec _ _ _ _ _ b, {intro x, exact is_trunc_of_le _ !minus_one_le_succ}, { refine @trunc.rec _ _ _ _ _ b, {intro x, exact is_trunc_of_le _ !minus_one_le_succ _ },
clear b, intro b, induction H b with a p, clear b, intro b, induction H b with a p,
exact tr (fiber.mk (tr a) (ap tr p))} exact tr (fiber.mk (tr a) (ap tr p))}
end end
@ -748,7 +748,7 @@ namespace trunc
is_trunc_trunc_of_is_trunc A n m is_trunc_trunc_of_is_trunc A n m
definition is_contr_ptrunc_minus_one (A : Type*) : is_contr (ptrunc -1 A) := definition is_contr_ptrunc_minus_one (A : Type*) : is_contr (ptrunc -1 A) :=
is_contr_of_inhabited_prop pt is_contr_of_inhabited_prop pt _
/- pointed maps involving ptrunc -/ /- pointed maps involving ptrunc -/
definition ptrunc_functor [constructor] {X Y : Type*} (n : ℕ₋₂) (f : X →* Y) definition ptrunc_functor [constructor] {X Y : Type*} (n : ℕ₋₂) (f : X →* Y)
@ -767,7 +767,7 @@ namespace trunc
definition ptrunc_functor_le {k l : ℕ₋₂} (p : l ≤ k) (X : Type*) definition ptrunc_functor_le {k l : ℕ₋₂} (p : l ≤ k) (X : Type*)
: ptrunc k X →* ptrunc l X := : ptrunc k X →* ptrunc l X :=
have is_trunc k (ptrunc l X), from is_trunc_of_le _ p, have is_trunc k (ptrunc l X), from is_trunc_of_le _ p _,
ptrunc.elim _ (ptr l X) ptrunc.elim _ (ptr l X)
/- pointed equivalences involving ptrunc -/ /- pointed equivalences involving ptrunc -/
@ -879,6 +879,14 @@ namespace trunc
{ esimp, refine !ap_con⁻¹ ⬝ _, exact ap02 tr !to_homotopy_pt}, { esimp, refine !ap_con⁻¹ ⬝ _, exact ap02 tr !to_homotopy_pt},
end end
definition ptrunc_functor_pconst [constructor] (n : ℕ₋₂) (X Y : Type*) :
ptrunc_functor n (pconst X Y) ~* pconst (ptrunc n X) (ptrunc n Y) :=
begin
fapply phomotopy.mk,
{ intro x, induction x with x, reflexivity },
{ reflexivity }
end
definition pcast_ptrunc [constructor] (n : ℕ₋₂) {A B : Type*} (p : A = B) : definition pcast_ptrunc [constructor] (n : ℕ₋₂) {A B : Type*} (p : A = B) :
pcast (ap (ptrunc n) p) ~* ptrunc_functor n (pcast p) := pcast (ap (ptrunc n) p) ~* ptrunc_functor n (pcast p) :=
begin begin
@ -887,6 +895,7 @@ namespace trunc
{ induction p, reflexivity} { induction p, reflexivity}
end end
definition ptrunc_elim_ptr [constructor] (n : ℕ₋₂) {X Y : Type*} [is_trunc n Y] (f : X →* Y) : definition ptrunc_elim_ptr [constructor] (n : ℕ₋₂) {X Y : Type*} [is_trunc n Y] (f : X →* Y) :
ptrunc.elim n f ∘* ptr n X ~* f := ptrunc.elim n f ∘* ptr n X ~* f :=
begin begin
@ -991,8 +1000,8 @@ namespace trunc
-- The following pointed equivalence can be defined more easily, but now we get the right maps definitionally -- The following pointed equivalence can be defined more easily, but now we get the right maps definitionally
definition ptrunc_pequiv_ptrunc_of_is_trunc {n m k : ℕ₋₂} {A : Type*} definition ptrunc_pequiv_ptrunc_of_is_trunc {n m k : ℕ₋₂} {A : Type*}
(H1 : n ≤ m) (H2 : n ≤ k) (H : is_trunc n A) : ptrunc m A ≃* ptrunc k A := (H1 : n ≤ m) (H2 : n ≤ k) (H : is_trunc n A) : ptrunc m A ≃* ptrunc k A :=
have is_trunc m A, from is_trunc_of_le A H1, have is_trunc m A, from is_trunc_of_le A H1 _,
have is_trunc k A, from is_trunc_of_le A H2, have is_trunc k A, from is_trunc_of_le A H2 _,
pequiv.MK (ptrunc.elim _ (ptr k A)) (ptrunc.elim _ (ptr m A)) pequiv.MK (ptrunc.elim _ (ptr k A)) (ptrunc.elim _ (ptr m A))
abstract begin abstract begin
refine !ptrunc_elim_pcompose⁻¹* ⬝* _, refine !ptrunc_elim_pcompose⁻¹* ⬝* _,
@ -1046,7 +1055,7 @@ namespace trunc
equiv.mk _ (is_embedding_ptrunctype_to_pType n X Y) ⬝e pType_eq_equiv X Y equiv.mk _ (is_embedding_ptrunctype_to_pType n X Y) ⬝e pType_eq_equiv X Y
definition Prop_eq {P Q : Prop} (H : P ↔ Q) : P = Q := definition Prop_eq {P Q : Prop} (H : P ↔ Q) : P = Q :=
tua (equiv_of_is_prop (iff.mp H) (iff.mpr H)) tua (equiv_of_is_prop (iff.mp H) (iff.mpr H) _ _)
end trunc open trunc end trunc open trunc
@ -1085,7 +1094,8 @@ namespace function
definition is_equiv_equiv_is_embedding_times_is_surjective [constructor] (f : A → B) definition is_equiv_equiv_is_embedding_times_is_surjective [constructor] (f : A → B)
: is_equiv f ≃ (is_embedding f × is_surjective f) := : is_equiv f ≃ (is_embedding f × is_surjective f) :=
equiv_of_is_prop (λH, (_, _)) equiv_of_is_prop (λH, (_, _))
(λP, prod.rec_on P (λH₁ H₂, !is_equiv_of_is_surjective_of_is_embedding)) (λP, prod.rec_on P (λH₁ H₂, !is_equiv_of_is_surjective_of_is_embedding))
_ _
/- /-
Theorem 8.8.1: Theorem 8.8.1:
@ -1119,7 +1129,7 @@ namespace function
note r := @(inj' (trunc_functor 0 f)) _ (tr a) (tr a') q, note r := @(inj' (trunc_functor 0 f)) _ (tr a) (tr a') q,
induction (tr_eq_tr_equiv _ _ _ r) with s, induction (tr_eq_tr_equiv _ _ _ r) with s,
induction s, induction s,
apply is_equiv.homotopy_closed (ap1 (pmap_of_map f a)), refine is_equiv.homotopy_closed (ap1 (pmap_of_map f a)) _ (H a),
intro p, apply idp_con intro p, apply idp_con
end, end,
is_equiv_of_is_surjective_trunc_of_is_embedding f is_equiv_of_is_surjective_trunc_of_is_embedding f
@ -1203,9 +1213,6 @@ namespace int
{ exact nat.zero_le m } { exact nat.zero_le m }
end end
definition not_neg_succ_le_of_nat {n m : } : ¬m ≤ -[1+n] :=
by cases m: exact id
definition maxm2_monotone {n m : } (H : n ≤ m) : maxm2 n ≤ maxm2 m := definition maxm2_monotone {n m : } (H : n ≤ m) : maxm2 n ≤ maxm2 m :=
begin begin
induction n with n n, induction n with n n,