style(*): rename is_hprop/is_hset to is_prop/is_set
This commit is contained in:
parent
816237315c
commit
4e2cc66061
70 changed files with 384 additions and 384 deletions
|
@ -116,7 +116,7 @@ namespace category
|
||||||
apply eq_of_fn_eq_fn !category.sigma_char,
|
apply eq_of_fn_eq_fn !category.sigma_char,
|
||||||
fapply sigma_eq,
|
fapply sigma_eq,
|
||||||
{ induction C, induction D, esimp, exact precategory_eq @p q},
|
{ induction C, induction D, esimp, exact precategory_eq @p q},
|
||||||
{ unfold is_univalent, apply is_hprop.elimo},
|
{ unfold is_univalent, apply is_prop.elimo},
|
||||||
end
|
end
|
||||||
|
|
||||||
definition category_eq_of_equiv {ob : Type}
|
definition category_eq_of_equiv {ob : Type}
|
||||||
|
|
|
@ -98,7 +98,7 @@ namespace category
|
||||||
{ intro u, exact (comma_morphism.mk u.1 u.2.1 u.2.2)},
|
{ intro u, exact (comma_morphism.mk u.1 u.2.1 u.2.2)},
|
||||||
{ intro f, cases f with g h p p', exact ⟨g, h, p⟩},
|
{ intro f, cases f with g h p p', exact ⟨g, h, p⟩},
|
||||||
{ intro f, cases f with g h p p', esimp,
|
{ intro f, cases f with g h p p', esimp,
|
||||||
apply ap (comma_morphism.mk' g h p), apply is_hprop.elim},
|
apply ap (comma_morphism.mk' g h p), apply is_prop.elim},
|
||||||
{ intro u, cases u with u1 u2, cases u2 with u2 u3, reflexivity},
|
{ intro u, cases u with u1 u2, cases u2 with u2 u3, reflexivity},
|
||||||
end
|
end
|
||||||
|
|
||||||
|
@ -113,8 +113,8 @@ namespace category
|
||||||
begin
|
begin
|
||||||
cases f with g h p₁ p₁', cases f' with g' h' p₂ p₂', cases p, cases q,
|
cases f with g h p₁ p₁', cases f' with g' h' p₂ p₂', cases p, cases q,
|
||||||
apply ap011 (comma_morphism.mk' g' h'),
|
apply ap011 (comma_morphism.mk' g' h'),
|
||||||
apply is_hprop.elim,
|
apply is_prop.elim,
|
||||||
apply is_hprop.elim
|
apply is_prop.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
definition comma_compose (g : comma_morphism y z) (f : comma_morphism x y) : comma_morphism x z :=
|
definition comma_compose (g : comma_morphism y z) (f : comma_morphism x y) : comma_morphism x z :=
|
||||||
|
|
|
@ -60,7 +60,7 @@ namespace category
|
||||||
theorem cone_hom_eq {f f' : cone_hom x y} (q : cone_to_hom f = cone_to_hom f') : f = f' :=
|
theorem cone_hom_eq {f f' : cone_hom x y} (q : cone_to_hom f = cone_to_hom f') : f = f' :=
|
||||||
begin
|
begin
|
||||||
induction f, induction f', esimp at *, induction q, apply ap (cone_hom.mk f),
|
induction f, induction f', esimp at *, induction q, apply ap (cone_hom.mk f),
|
||||||
apply @is_hprop.elim, apply pi.is_trunc_pi, intro x, apply is_trunc_eq, -- type class fails
|
apply @is_prop.elim, apply pi.is_trunc_pi, intro x, apply is_trunc_eq, -- type class fails
|
||||||
end
|
end
|
||||||
|
|
||||||
variable (F)
|
variable (F)
|
||||||
|
@ -120,7 +120,7 @@ namespace category
|
||||||
{ intro v, exact cone_iso.mk v.1 v.2},
|
{ intro v, exact cone_iso.mk v.1 v.2},
|
||||||
{ intro v, induction v with f p, fapply sigma_eq: esimp,
|
{ intro v, induction v with f p, fapply sigma_eq: esimp,
|
||||||
{ apply iso_eq, reflexivity},
|
{ apply iso_eq, reflexivity},
|
||||||
{ apply is_hprop.elimo, apply is_trunc_pi, intro i, apply is_hprop_hom_eq}},
|
{ apply is_prop.elimo, apply is_trunc_pi, intro i, apply is_prop_hom_eq}},
|
||||||
{ intro h, esimp, apply iso_eq, apply cone_hom_eq, reflexivity},
|
{ intro h, esimp, apply iso_eq, apply cone_hom_eq, reflexivity},
|
||||||
end
|
end
|
||||||
|
|
||||||
|
@ -134,8 +134,8 @@ namespace category
|
||||||
{ intro v, induction v with p q, induction x with c η, induction y with c' η', esimp at *,
|
{ intro v, induction v with p q, induction x with c η, induction y with c' η', esimp at *,
|
||||||
induction p, esimp, fapply sigma_eq: esimp,
|
induction p, esimp, fapply sigma_eq: esimp,
|
||||||
{ apply c_cone_obj_eq},
|
{ apply c_cone_obj_eq},
|
||||||
{ apply is_hprop.elimo, apply is_trunc_pi, intro i, apply is_hprop_hom_eq}},
|
{ apply is_prop.elimo, apply is_trunc_pi, intro i, apply is_prop_hom_eq}},
|
||||||
{ intro r, induction r, esimp, induction x, esimp, apply ap02, apply is_hprop.elim},
|
{ intro r, induction r, esimp, induction x, esimp, apply ap02, apply is_prop.elim},
|
||||||
end
|
end
|
||||||
|
|
||||||
section is_univalent
|
section is_univalent
|
||||||
|
|
|
@ -31,16 +31,16 @@ namespace category
|
||||||
definition Groupoid_of_1_type [constructor] (A : Type) [H : is_trunc 1 A] : Groupoid :=
|
definition Groupoid_of_1_type [constructor] (A : Type) [H : is_trunc 1 A] : Groupoid :=
|
||||||
groupoid.Mk _ (groupoid_of_1_type A)
|
groupoid.Mk _ (groupoid_of_1_type A)
|
||||||
|
|
||||||
definition discrete_precategory [constructor] (A : Type) [H : is_hset A] : precategory A :=
|
definition discrete_precategory [constructor] (A : Type) [H : is_set A] : precategory A :=
|
||||||
precategory_of_1_type A
|
precategory_of_1_type A
|
||||||
|
|
||||||
definition discrete_groupoid [constructor] (A : Type) [H : is_hset A] : groupoid A :=
|
definition discrete_groupoid [constructor] (A : Type) [H : is_set A] : groupoid A :=
|
||||||
groupoid_of_1_type A
|
groupoid_of_1_type A
|
||||||
|
|
||||||
definition Discrete_precategory [constructor] (A : Type) [H : is_hset A] : Precategory :=
|
definition Discrete_precategory [constructor] (A : Type) [H : is_set A] : Precategory :=
|
||||||
precategory.Mk (discrete_precategory A)
|
precategory.Mk (discrete_precategory A)
|
||||||
|
|
||||||
definition Discrete_groupoid [constructor] (A : Type) [H : is_hset A] : Groupoid :=
|
definition Discrete_groupoid [constructor] (A : Type) [H : is_set A] : Groupoid :=
|
||||||
groupoid.Mk _ (discrete_groupoid A)
|
groupoid.Mk _ (discrete_groupoid A)
|
||||||
|
|
||||||
definition c2 [constructor] : Precategory := Discrete_precategory bool
|
definition c2 [constructor] : Precategory := Discrete_precategory bool
|
||||||
|
|
|
@ -14,7 +14,7 @@ open bool unit is_trunc sum eq functor equiv
|
||||||
namespace category
|
namespace category
|
||||||
|
|
||||||
variables {A : Type} (R : A → A → Type) (H : Π⦃a b c⦄, R a b → R b c → empty)
|
variables {A : Type} (R : A → A → Type) (H : Π⦃a b c⦄, R a b → R b c → empty)
|
||||||
[HR : Πa b, is_hset (R a b)] [HA : is_trunc 1 A]
|
[HR : Πa b, is_set (R a b)] [HA : is_trunc 1 A]
|
||||||
|
|
||||||
include H HR HA
|
include H HR HA
|
||||||
|
|
||||||
|
@ -63,7 +63,7 @@ namespace category
|
||||||
| f2 : equalizer_category_hom ff tt
|
| f2 : equalizer_category_hom ff tt
|
||||||
|
|
||||||
open equalizer_category_hom
|
open equalizer_category_hom
|
||||||
theorem is_hset_equalizer_category_hom (b₁ b₂ : bool) : is_hset (equalizer_category_hom b₁ b₂) :=
|
theorem is_set_equalizer_category_hom (b₁ b₂ : bool) : is_set (equalizer_category_hom b₁ b₂) :=
|
||||||
begin
|
begin
|
||||||
assert H : Πb b', equalizer_category_hom b b' ≃ bool.rec (bool.rec empty bool) (λb, empty) b b',
|
assert H : Πb b', equalizer_category_hom b b' ≃ bool.rec (bool.rec empty bool) (λb, empty) b b',
|
||||||
{ intro b b', fapply equiv.MK,
|
{ intro b b', fapply equiv.MK,
|
||||||
|
@ -75,7 +75,7 @@ namespace category
|
||||||
induction b₁: induction b₂: exact _
|
induction b₁: induction b₂: exact _
|
||||||
end
|
end
|
||||||
|
|
||||||
local attribute is_hset_equalizer_category_hom [instance]
|
local attribute is_set_equalizer_category_hom [instance]
|
||||||
definition equalizer_category [constructor] : Precategory :=
|
definition equalizer_category [constructor] : Precategory :=
|
||||||
sparse_category
|
sparse_category
|
||||||
equalizer_category_hom
|
equalizer_category_hom
|
||||||
|
@ -107,8 +107,8 @@ namespace category
|
||||||
| f2 : pullback_category_hom BL BR
|
| f2 : pullback_category_hom BL BR
|
||||||
|
|
||||||
open pullback_category_hom
|
open pullback_category_hom
|
||||||
theorem is_hset_pullback_category_hom (b₁ b₂ : pullback_category_ob)
|
theorem is_set_pullback_category_hom (b₁ b₂ : pullback_category_ob)
|
||||||
: is_hset (pullback_category_hom b₁ b₂) :=
|
: is_set (pullback_category_hom b₁ b₂) :=
|
||||||
begin
|
begin
|
||||||
assert H : Πb b', pullback_category_hom b b' ≃
|
assert H : Πb b', pullback_category_hom b b' ≃
|
||||||
pullback_category_ob.rec (λb, empty) (λb, empty)
|
pullback_category_ob.rec (λb, empty) (λb, empty)
|
||||||
|
@ -122,7 +122,7 @@ namespace category
|
||||||
induction b₁: induction b₂: exact _
|
induction b₁: induction b₂: exact _
|
||||||
end
|
end
|
||||||
|
|
||||||
local attribute is_hset_pullback_category_hom pullback_category_ob_decidable_equality [instance]
|
local attribute is_set_pullback_category_hom pullback_category_ob_decidable_equality [instance]
|
||||||
definition pullback_category [constructor] : Precategory :=
|
definition pullback_category [constructor] : Precategory :=
|
||||||
sparse_category
|
sparse_category
|
||||||
pullback_category_hom
|
pullback_category_hom
|
||||||
|
|
|
@ -47,7 +47,7 @@ namespace functor
|
||||||
fapply (apd011 nat_trans.mk),
|
fapply (apd011 nat_trans.mk),
|
||||||
apply eq_of_homotopy, intro c, apply left_inverse,
|
apply eq_of_homotopy, intro c, apply left_inverse,
|
||||||
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
|
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
|
||||||
apply is_hset.elim
|
apply is_set.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
definition nat_trans_right_inverse : η ∘n nat_trans_inverse η = 1 :=
|
definition nat_trans_right_inverse : η ∘n nat_trans_inverse η = 1 :=
|
||||||
|
@ -55,7 +55,7 @@ namespace functor
|
||||||
fapply (apd011 nat_trans.mk),
|
fapply (apd011 nat_trans.mk),
|
||||||
apply eq_of_homotopy, intro c, apply right_inverse,
|
apply eq_of_homotopy, intro c, apply right_inverse,
|
||||||
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
|
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
|
||||||
apply is_hset.elim
|
apply is_set.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_natural_iso [constructor] : is_iso η :=
|
definition is_natural_iso [constructor] : is_iso η :=
|
||||||
|
|
|
@ -38,10 +38,10 @@ namespace category
|
||||||
|
|
||||||
definition initial_functor_op (C : Precategory)
|
definition initial_functor_op (C : Precategory)
|
||||||
: (initial_functor C)ᵒᵖᶠ = initial_functor Cᵒᵖ :=
|
: (initial_functor C)ᵒᵖᶠ = initial_functor Cᵒᵖ :=
|
||||||
by apply @is_hprop.elim (0 ⇒ Cᵒᵖ)
|
by apply @is_prop.elim (0 ⇒ Cᵒᵖ)
|
||||||
|
|
||||||
definition initial_functor_comp {C D : Precategory} (F : C ⇒ D)
|
definition initial_functor_comp {C D : Precategory} (F : C ⇒ D)
|
||||||
: F ∘f initial_functor C = initial_functor D :=
|
: F ∘f initial_functor C = initial_functor D :=
|
||||||
by apply @is_hprop.elim (0 ⇒ D)
|
by apply @is_prop.elim (0 ⇒ D)
|
||||||
|
|
||||||
end category
|
end category
|
||||||
|
|
|
@ -21,7 +21,7 @@ namespace category
|
||||||
(λ a b f, !id_right)
|
(λ a b f, !id_right)
|
||||||
(λ a b f, !id_left)
|
(λ a b f, !id_left)
|
||||||
(λ a, !id_id)
|
(λ a, !id_id)
|
||||||
(λ a b, !is_hset_hom)
|
(λ a b, !is_set_hom)
|
||||||
|
|
||||||
definition Opposite [reducible] [constructor] (C : Precategory) : Precategory :=
|
definition Opposite [reducible] [constructor] (C : Precategory) : Precategory :=
|
||||||
precategory.Mk (opposite C)
|
precategory.Mk (opposite C)
|
||||||
|
|
|
@ -19,12 +19,12 @@ namespace category
|
||||||
sum.rec (λc, sum.rec (λc', lift (c ⟶ c')) (λd, lift empty))
|
sum.rec (λc, sum.rec (λc', lift (c ⟶ c')) (λd, lift empty))
|
||||||
(λd, sum.rec (λc, lift empty) (λd', lift (d ⟶ d')))
|
(λd, sum.rec (λc, lift empty) (λd', lift (d ⟶ d')))
|
||||||
|
|
||||||
theorem is_hset_sum_hom {obC : Type} {obD : Type}
|
theorem is_set_sum_hom {obC : Type} {obD : Type}
|
||||||
(C : precategory obC) (D : precategory obD) (x y : obC + obD)
|
(C : precategory obC) (D : precategory obD) (x y : obC + obD)
|
||||||
: is_hset (sum_hom C D x y) :=
|
: is_set (sum_hom C D x y) :=
|
||||||
by induction x: induction y: esimp at *: exact _
|
by induction x: induction y: esimp at *: exact _
|
||||||
|
|
||||||
local attribute is_hset_sum_hom [instance]
|
local attribute is_set_sum_hom [instance]
|
||||||
|
|
||||||
definition precategory_sum [constructor] [instance] (obC obD : Type)
|
definition precategory_sum [constructor] [instance] (obC obD : Type)
|
||||||
[C : precategory obC] [D : precategory obD] : precategory (obC + obD) :=
|
[C : precategory obC] [D : precategory obD] : precategory (obC + obD) :=
|
||||||
|
|
|
@ -32,8 +32,8 @@ namespace category
|
||||||
is_contr.mk (terminal_functor C)
|
is_contr.mk (terminal_functor C)
|
||||||
begin
|
begin
|
||||||
intro F, fapply functor_eq,
|
intro F, fapply functor_eq,
|
||||||
{ intro x, apply @is_hprop.elim unit},
|
{ intro x, apply @is_prop.elim unit},
|
||||||
{ intro x y f, apply @is_hprop.elim unit}
|
{ intro x y f, apply @is_prop.elim unit}
|
||||||
end
|
end
|
||||||
|
|
||||||
definition terminal_functor_op (C : Precategory)
|
definition terminal_functor_op (C : Precategory)
|
||||||
|
|
|
@ -43,15 +43,15 @@ namespace category
|
||||||
abbreviation counit_unit_eq [unfold 4] := @is_left_adjoint.H
|
abbreviation counit_unit_eq [unfold 4] := @is_left_adjoint.H
|
||||||
abbreviation unit_counit_eq [unfold 4] := @is_left_adjoint.K
|
abbreviation unit_counit_eq [unfold 4] := @is_left_adjoint.K
|
||||||
|
|
||||||
theorem is_hprop_is_left_adjoint [instance] {C : Category} {D : Precategory} (F : C ⇒ D)
|
theorem is_prop_is_left_adjoint [instance] {C : Category} {D : Precategory} (F : C ⇒ D)
|
||||||
: is_hprop (is_left_adjoint F) :=
|
: is_prop (is_left_adjoint F) :=
|
||||||
begin
|
begin
|
||||||
apply is_hprop.mk,
|
apply is_prop.mk,
|
||||||
intro G G', cases G with G η ε H K, cases G' with G' η' ε' H' K',
|
intro G G', cases G with G η ε H K, cases G' with G' η' ε' H' K',
|
||||||
assert lem₁ : Π(p : G = G'), p ▸ η = η' → p ▸ ε = ε'
|
assert lem₁ : Π(p : G = G'), p ▸ η = η' → p ▸ ε = ε'
|
||||||
→ is_left_adjoint.mk G η ε H K = is_left_adjoint.mk G' η' ε' H' K',
|
→ is_left_adjoint.mk G η ε H K = is_left_adjoint.mk G' η' ε' H' K',
|
||||||
{ intros p q r, induction p, induction q, induction r, esimp,
|
{ intros p q r, induction p, induction q, induction r, esimp,
|
||||||
apply apd011 (is_left_adjoint.mk G η ε) !is_hprop.elim !is_hprop.elim},
|
apply apd011 (is_left_adjoint.mk G η ε) !is_prop.elim !is_prop.elim},
|
||||||
assert lem₂ : Π (d : carrier D),
|
assert lem₂ : Π (d : carrier D),
|
||||||
(to_fun_hom G (natural_map ε' d) ∘
|
(to_fun_hom G (natural_map ε' d) ∘
|
||||||
natural_map η (to_fun_ob G' d)) ∘
|
natural_map η (to_fun_ob G' d)) ∘
|
||||||
|
|
|
@ -117,24 +117,24 @@ namespace category
|
||||||
apply is_equiv_of_is_surjective_of_is_embedding,
|
apply is_equiv_of_is_surjective_of_is_embedding,
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem is_hprop_fully_faithful [instance] (F : C ⇒ D) : is_hprop (fully_faithful F) :=
|
theorem is_prop_fully_faithful [instance] (F : C ⇒ D) : is_prop (fully_faithful F) :=
|
||||||
by unfold fully_faithful; exact _
|
by unfold fully_faithful; exact _
|
||||||
|
|
||||||
theorem is_hprop_full [instance] (F : C ⇒ D) : is_hprop (full F) :=
|
theorem is_prop_full [instance] (F : C ⇒ D) : is_prop (full F) :=
|
||||||
by unfold full; exact _
|
by unfold full; exact _
|
||||||
|
|
||||||
theorem is_hprop_faithful [instance] (F : C ⇒ D) : is_hprop (faithful F) :=
|
theorem is_prop_faithful [instance] (F : C ⇒ D) : is_prop (faithful F) :=
|
||||||
by unfold faithful; exact _
|
by unfold faithful; exact _
|
||||||
|
|
||||||
theorem is_hprop_essentially_surjective [instance] (F : C ⇒ D)
|
theorem is_prop_essentially_surjective [instance] (F : C ⇒ D)
|
||||||
: is_hprop (essentially_surjective F) :=
|
: is_prop (essentially_surjective F) :=
|
||||||
by unfold essentially_surjective; exact _
|
by unfold essentially_surjective; exact _
|
||||||
|
|
||||||
theorem is_hprop_is_weak_equivalence [instance] (F : C ⇒ D) : is_hprop (is_weak_equivalence F) :=
|
theorem is_prop_is_weak_equivalence [instance] (F : C ⇒ D) : is_prop (is_weak_equivalence F) :=
|
||||||
by unfold is_weak_equivalence; exact _
|
by unfold is_weak_equivalence; exact _
|
||||||
|
|
||||||
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_hprop (λ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
|
||||||
|
|
|
@ -56,7 +56,7 @@ namespace functor
|
||||||
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} (id₁ id₂ comp₁ comp₂)
|
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} (id₁ id₂ comp₁ comp₂)
|
||||||
(pF : F₁ = F₂) (pH : pF ▸ H₁ = H₂)
|
(pF : F₁ = F₂) (pH : pF ▸ H₁ = H₂)
|
||||||
: functor.mk F₁ H₁ id₁ comp₁ = functor.mk F₂ H₂ id₂ comp₂ :=
|
: functor.mk F₁ H₁ id₁ comp₁ = functor.mk F₂ H₂ id₂ comp₂ :=
|
||||||
apd01111 functor.mk pF pH !is_hprop.elim !is_hprop.elim
|
apd01111 functor.mk pF pH !is_prop.elim !is_prop.elim
|
||||||
|
|
||||||
definition functor_eq' {F₁ F₂ : C ⇒ D} : Π(p : to_fun_ob F₁ = to_fun_ob F₂),
|
definition functor_eq' {F₁ F₂ : C ⇒ D} : Π(p : to_fun_ob F₁ = to_fun_ob F₂),
|
||||||
(transport (λx, Πa b f, hom (x a) (x b)) p @(to_fun_hom F₁) = @(to_fun_hom F₂)) → F₁ = F₂ :=
|
(transport (λx, Πa b f, hom (x a) (x b)) p @(to_fun_hom F₁) = @(to_fun_hom F₂)) → F₁ = F₂ :=
|
||||||
|
@ -181,10 +181,10 @@ namespace functor
|
||||||
q (respect_comp F g f) end qed
|
q (respect_comp F g f) end qed
|
||||||
|
|
||||||
section
|
section
|
||||||
local attribute precategory.is_hset_hom [instance] [priority 1001]
|
local attribute precategory.is_set_hom [instance] [priority 1001]
|
||||||
local attribute trunctype.struct [instance] [priority 1] -- remove after #842 is closed
|
local attribute trunctype.struct [instance] [priority 1] -- remove after #842 is closed
|
||||||
protected theorem is_hset_functor [instance]
|
protected theorem is_set_functor [instance]
|
||||||
[HD : is_hset D] : is_hset (functor C D) :=
|
[HD : is_set D] : is_set (functor C D) :=
|
||||||
by apply is_trunc_equiv_closed; apply functor.sigma_char
|
by apply is_trunc_equiv_closed; apply functor.sigma_char
|
||||||
end
|
end
|
||||||
|
|
||||||
|
@ -193,8 +193,8 @@ namespace functor
|
||||||
(id comp) : functor_mk_eq' id id comp comp (idpath F) (idpath H) = idp :=
|
(id comp) : functor_mk_eq' id id comp comp (idpath F) (idpath H) = idp :=
|
||||||
begin
|
begin
|
||||||
fapply apd011 (apd01111 functor.mk idp idp),
|
fapply apd011 (apd01111 functor.mk idp idp),
|
||||||
apply is_hset.elim,
|
apply is_set.elim,
|
||||||
apply is_hset.elim
|
apply is_set.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
definition functor_eq'_idp (F : C ⇒ D) : functor_eq' idp idp = (idpath F) :=
|
definition functor_eq'_idp (F : C ⇒ D) : functor_eq' idp idp = (idpath F) :=
|
||||||
|
@ -210,7 +210,7 @@ namespace functor
|
||||||
|
|
||||||
theorem functor_eq2' {F₁ F₂ : C ⇒ D} {p₁ p₂ : to_fun_ob F₁ = to_fun_ob F₂} (q₁ q₂)
|
theorem functor_eq2' {F₁ F₂ : C ⇒ D} {p₁ p₂ : to_fun_ob F₁ = to_fun_ob F₂} (q₁ q₂)
|
||||||
(r : p₁ = p₂) : functor_eq' p₁ q₁ = functor_eq' p₂ q₂ :=
|
(r : p₁ = p₂) : functor_eq' p₁ q₁ = functor_eq' p₂ q₂ :=
|
||||||
by cases r; apply (ap (functor_eq' p₂)); apply is_hprop.elim
|
by cases r; apply (ap (functor_eq' p₂)); apply is_prop.elim
|
||||||
|
|
||||||
theorem functor_eq2 {F₁ F₂ : C ⇒ D} (p q : F₁ = F₂) (r : ap010 to_fun_ob p ~ ap010 to_fun_ob q)
|
theorem functor_eq2 {F₁ F₂ : C ⇒ D} (p q : F₁ = F₂) (r : ap010 to_fun_ob p ~ ap010 to_fun_ob q)
|
||||||
: p = q :=
|
: p = q :=
|
||||||
|
|
|
@ -228,8 +228,8 @@ namespace category
|
||||||
definition equivalence_of_isomorphism [constructor] (F : C ≅c D) : C ≃c D :=
|
definition equivalence_of_isomorphism [constructor] (F : C ≅c D) : C ≃c D :=
|
||||||
equivalence.mk F _
|
equivalence.mk F _
|
||||||
|
|
||||||
theorem is_hprop_is_equivalence [instance] {C : Category} {D : Precategory} (F : C ⇒ D)
|
theorem is_prop_is_equivalence [instance] {C : Category} {D : Precategory} (F : C ⇒ D)
|
||||||
: is_hprop (is_equivalence F) :=
|
: is_prop (is_equivalence F) :=
|
||||||
begin
|
begin
|
||||||
assert f : is_equivalence F ≃ Σ(H : is_left_adjoint F), is_iso (unit F) × is_iso (counit F),
|
assert f : is_equivalence F ≃ Σ(H : is_left_adjoint F), is_iso (unit F) × is_iso (counit F),
|
||||||
{ fapply equiv.MK,
|
{ fapply equiv.MK,
|
||||||
|
@ -241,7 +241,7 @@ namespace category
|
||||||
apply is_trunc_equiv_closed_rev, exact f,
|
apply is_trunc_equiv_closed_rev, exact f,
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem is_hprop_is_isomorphism [instance] (F : C ⇒ D) : is_hprop (is_isomorphism F) :=
|
theorem is_prop_is_isomorphism [instance] (F : C ⇒ D) : is_prop (is_isomorphism F) :=
|
||||||
by unfold is_isomorphism; exact _
|
by unfold is_isomorphism; exact _
|
||||||
|
|
||||||
/- closure properties -/
|
/- closure properties -/
|
||||||
|
@ -332,13 +332,13 @@ namespace category
|
||||||
definition equivalence_eq {C : Category} {D : Precategory} {F F' : C ≃c D}
|
definition equivalence_eq {C : Category} {D : Precategory} {F F' : C ≃c D}
|
||||||
(p : equivalence.to_functor F = equivalence.to_functor F') : F = F' :=
|
(p : equivalence.to_functor F = equivalence.to_functor F') : F = F' :=
|
||||||
begin
|
begin
|
||||||
induction F, induction F', exact apd011 equivalence.mk p !is_hprop.elim
|
induction F, induction F', exact apd011 equivalence.mk p !is_prop.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
definition isomorphism_eq {F F' : C ≅c D}
|
definition isomorphism_eq {F F' : C ≅c D}
|
||||||
(p : isomorphism.to_functor F = isomorphism.to_functor F') : F = F' :=
|
(p : isomorphism.to_functor F = isomorphism.to_functor F') : F = F' :=
|
||||||
begin
|
begin
|
||||||
induction F, induction F', exact apd011 isomorphism.mk p !is_hprop.elim
|
induction F, induction F', exact apd011 isomorphism.mk p !is_prop.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_equiv_isomorphism_of_equivalence [constructor] (C D : Category)
|
definition is_equiv_isomorphism_of_equivalence [constructor] (C D : Category)
|
||||||
|
|
|
@ -138,12 +138,12 @@ namespace yoneda
|
||||||
|
|
||||||
section
|
section
|
||||||
set_option apply.class_instance false
|
set_option apply.class_instance false
|
||||||
definition is_hprop_representable {C : Category} (F : Cᵒᵖ ⇒ cset)
|
definition is_prop_representable {C : Category} (F : Cᵒᵖ ⇒ cset)
|
||||||
: is_hprop (is_representable F) :=
|
: is_prop (is_representable F) :=
|
||||||
begin
|
begin
|
||||||
fapply is_trunc_equiv_closed,
|
fapply is_trunc_equiv_closed,
|
||||||
{ exact proof fiber.sigma_char ɏ F qed ⬝e sigma.sigma_equiv_sigma_id (λc, !eq_equiv_iso)},
|
{ exact proof fiber.sigma_char ɏ F qed ⬝e sigma.sigma_equiv_sigma_id (λc, !eq_equiv_iso)},
|
||||||
{ apply function.is_hprop_fiber_of_is_embedding, apply is_embedding_yoneda_embedding}
|
{ apply function.is_prop_fiber_of_is_embedding, apply is_embedding_yoneda_embedding}
|
||||||
end
|
end
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -27,7 +27,7 @@ namespace category
|
||||||
begin
|
begin
|
||||||
fapply groupoid.mk, fapply precategory.mk,
|
fapply groupoid.mk, fapply precategory.mk,
|
||||||
intros, exact A,
|
intros, exact A,
|
||||||
intros, apply (@group.is_hset_carrier A G),
|
intros, apply (@group.is_set_carrier A G),
|
||||||
intros [a, b, c, g, h], exact (@group.mul A G g h),
|
intros [a, b, c, g, h], exact (@group.mul A G g h),
|
||||||
intro a, exact (@group.one A G),
|
intro a, exact (@group.one A G),
|
||||||
intros, exact (@group.mul_assoc A G h g f)⁻¹,
|
intros, exact (@group.mul_assoc A G h g f)⁻¹,
|
||||||
|
@ -43,7 +43,7 @@ namespace category
|
||||||
begin
|
begin
|
||||||
fapply group.mk,
|
fapply group.mk,
|
||||||
intro f g, apply (comp f g),
|
intro f g, apply (comp f g),
|
||||||
apply is_hset_hom,
|
apply is_set_hom,
|
||||||
intros f g h, apply (assoc f g h)⁻¹,
|
intros f g h, apply (assoc f g h)⁻¹,
|
||||||
apply (ID a),
|
apply (ID a),
|
||||||
intro f, apply id_left,
|
intro f, apply id_left,
|
||||||
|
|
|
@ -114,16 +114,16 @@ namespace iso
|
||||||
[Hf : is_iso f] [Hg : is_iso g] : is_iso (g ∘ f) :=
|
[Hf : is_iso f] [Hg : is_iso g] : is_iso (g ∘ f) :=
|
||||||
!is_iso_of_split_epi_of_split_mono
|
!is_iso_of_split_epi_of_split_mono
|
||||||
|
|
||||||
theorem is_hprop_is_iso [instance] (f : hom a b) : is_hprop (is_iso f) :=
|
theorem is_prop_is_iso [instance] (f : hom a b) : is_prop (is_iso f) :=
|
||||||
begin
|
begin
|
||||||
apply is_hprop.mk, intro H H',
|
apply is_prop.mk, intro H H',
|
||||||
cases H with g li ri, cases H' with g' li' ri',
|
cases H with g li ri, cases H' with g' li' ri',
|
||||||
fapply (apd0111 (@is_iso.mk ob C a b f)),
|
fapply (apd0111 (@is_iso.mk ob C a b f)),
|
||||||
apply left_inverse_eq_right_inverse,
|
apply left_inverse_eq_right_inverse,
|
||||||
apply li,
|
apply li,
|
||||||
apply ri',
|
apply ri',
|
||||||
apply is_hprop.elim,
|
apply is_prop.elim,
|
||||||
apply is_hprop.elim,
|
apply is_prop.elim,
|
||||||
end
|
end
|
||||||
end iso open iso
|
end iso open iso
|
||||||
|
|
||||||
|
@ -175,7 +175,7 @@ namespace iso
|
||||||
|
|
||||||
definition iso_mk_eq {f f' : a ⟶ b} [H : is_iso f] [H' : is_iso f'] (p : f = f')
|
definition iso_mk_eq {f f' : a ⟶ b} [H : is_iso f] [H' : is_iso f'] (p : f = f')
|
||||||
: iso.mk f _ = iso.mk f' _ :=
|
: iso.mk f _ = iso.mk f' _ :=
|
||||||
apd011 iso.mk p !is_hprop.elim
|
apd011 iso.mk p !is_prop.elim
|
||||||
|
|
||||||
variable {C}
|
variable {C}
|
||||||
definition iso_eq {f f' : a ≅ b} (p : to_hom f = to_hom f') : f = f' :=
|
definition iso_eq {f f' : a ≅ b} (p : to_hom f = to_hom f') : f = f' :=
|
||||||
|
@ -194,7 +194,7 @@ namespace iso
|
||||||
end
|
end
|
||||||
|
|
||||||
-- The type of isomorphisms between two objects is a set
|
-- The type of isomorphisms between two objects is a set
|
||||||
definition is_hset_iso [instance] : is_hset (a ≅ b) :=
|
definition is_set_iso [instance] : is_set (a ≅ b) :=
|
||||||
begin
|
begin
|
||||||
apply is_trunc_is_equiv_closed,
|
apply is_trunc_is_equiv_closed,
|
||||||
apply equiv.to_is_equiv (!iso.sigma_char),
|
apply equiv.to_is_equiv (!iso.sigma_char),
|
||||||
|
|
|
@ -29,15 +29,15 @@ namespace category
|
||||||
!center
|
!center
|
||||||
|
|
||||||
definition hom_initial_eq [H : is_initial c'] (f f' : c' ⟶ c) : f = f' :=
|
definition hom_initial_eq [H : is_initial c'] (f f' : c' ⟶ c) : f = f' :=
|
||||||
!is_hprop.elim
|
!is_prop.elim
|
||||||
|
|
||||||
definition eq_initial_morphism [H : is_initial c'] (f : c' ⟶ c) : f = initial_morphism c c' :=
|
definition eq_initial_morphism [H : is_initial c'] (f : c' ⟶ c) : f = initial_morphism c c' :=
|
||||||
!is_hprop.elim
|
!is_prop.elim
|
||||||
|
|
||||||
definition initial_iso_initial {c c' : ob} (H : is_initial c) (K : is_initial c') : c ≅ c' :=
|
definition initial_iso_initial {c c' : ob} (H : is_initial c) (K : is_initial c') : c ≅ c' :=
|
||||||
iso_of_opposite_iso (@terminal_iso_terminal _ (opposite C) _ _ H K)
|
iso_of_opposite_iso (@terminal_iso_terminal _ (opposite C) _ _ H K)
|
||||||
|
|
||||||
theorem is_hprop_is_initial [instance] : is_hprop (is_initial c) := _
|
theorem is_prop_is_initial [instance] : is_prop (is_initial c) := _
|
||||||
|
|
||||||
omit C
|
omit C
|
||||||
|
|
||||||
|
@ -56,9 +56,9 @@ namespace category
|
||||||
initial_iso_initial (@has_initial_object.is_initial D H₁) (@has_initial_object.is_initial D H₂)
|
initial_iso_initial (@has_initial_object.is_initial D H₁) (@has_initial_object.is_initial D H₂)
|
||||||
|
|
||||||
set_option pp.coercions true
|
set_option pp.coercions true
|
||||||
theorem is_hprop_has_initial_object [instance] (D : Category)
|
theorem is_prop_has_initial_object [instance] (D : Category)
|
||||||
: is_hprop (has_initial_object D) :=
|
: is_prop (has_initial_object D) :=
|
||||||
is_hprop_has_terminal_object (Category_opposite D)
|
is_prop_has_terminal_object (Category_opposite D)
|
||||||
|
|
||||||
variable (D)
|
variable (D)
|
||||||
abbreviation has_colimits_of_shape := has_limits_of_shape Dᵒᵖ Iᵒᵖ
|
abbreviation has_colimits_of_shape := has_limits_of_shape Dᵒᵖ Iᵒᵖ
|
||||||
|
@ -76,12 +76,12 @@ namespace category
|
||||||
|
|
||||||
section
|
section
|
||||||
open pi
|
open pi
|
||||||
theorem is_hprop_has_colimits_of_shape [instance] (D : Category) (I : Precategory)
|
theorem is_prop_has_colimits_of_shape [instance] (D : Category) (I : Precategory)
|
||||||
: is_hprop (has_colimits_of_shape D I) :=
|
: is_prop (has_colimits_of_shape D I) :=
|
||||||
is_hprop_has_limits_of_shape (Category_opposite D) _
|
is_prop_has_limits_of_shape (Category_opposite D) _
|
||||||
|
|
||||||
theorem is_hprop_is_cocomplete [instance] (D : Category) : is_hprop (is_cocomplete D) :=
|
theorem is_prop_is_cocomplete [instance] (D : Category) : is_prop (is_cocomplete D) :=
|
||||||
is_hprop_is_complete (Category_opposite D)
|
is_prop_is_complete (Category_opposite D)
|
||||||
end
|
end
|
||||||
|
|
||||||
variables {D I} (F : I ⇒ D) [H : has_colimits_of_shape D I] {i j : I}
|
variables {D I} (F : I ⇒ D) [H : has_colimits_of_shape D I] {i j : I}
|
||||||
|
|
|
@ -26,17 +26,17 @@ namespace category
|
||||||
!center
|
!center
|
||||||
|
|
||||||
definition hom_terminal_eq [H : is_terminal c'] (f f' : c ⟶ c') : f = f' :=
|
definition hom_terminal_eq [H : is_terminal c'] (f f' : c ⟶ c') : f = f' :=
|
||||||
!is_hprop.elim
|
!is_prop.elim
|
||||||
|
|
||||||
definition eq_terminal_morphism [H : is_terminal c'] (f : c ⟶ c') : f = terminal_morphism c c' :=
|
definition eq_terminal_morphism [H : is_terminal c'] (f : c ⟶ c') : f = terminal_morphism c c' :=
|
||||||
!is_hprop.elim
|
!is_prop.elim
|
||||||
|
|
||||||
definition terminal_iso_terminal (c c' : ob) [H : is_terminal c] [K : is_terminal c']
|
definition terminal_iso_terminal (c c' : ob) [H : is_terminal c] [K : is_terminal c']
|
||||||
: c ≅ c' :=
|
: c ≅ c' :=
|
||||||
iso.MK !terminal_morphism !terminal_morphism !hom_terminal_eq !hom_terminal_eq
|
iso.MK !terminal_morphism !terminal_morphism !hom_terminal_eq !hom_terminal_eq
|
||||||
|
|
||||||
local attribute is_terminal [reducible]
|
local attribute is_terminal [reducible]
|
||||||
theorem is_hprop_is_terminal [instance] : is_hprop (is_terminal c) :=
|
theorem is_prop_is_terminal [instance] : is_prop (is_terminal c) :=
|
||||||
_
|
_
|
||||||
|
|
||||||
omit C
|
omit C
|
||||||
|
@ -53,13 +53,13 @@ namespace category
|
||||||
: @terminal_object D H₁ ≅ @terminal_object D H₂ :=
|
: @terminal_object D H₁ ≅ @terminal_object D H₂ :=
|
||||||
!terminal_iso_terminal
|
!terminal_iso_terminal
|
||||||
|
|
||||||
theorem is_hprop_has_terminal_object [instance] (D : Category)
|
theorem is_prop_has_terminal_object [instance] (D : Category)
|
||||||
: is_hprop (has_terminal_object D) :=
|
: is_prop (has_terminal_object D) :=
|
||||||
begin
|
begin
|
||||||
apply is_hprop.mk, intro t₁ t₂, induction t₁ with d₁ H₁, induction t₂ with d₂ H₂,
|
apply is_prop.mk, intro t₁ t₂, induction t₁ with d₁ H₁, induction t₂ with d₂ H₂,
|
||||||
assert p : d₁ = d₂,
|
assert p : d₁ = d₂,
|
||||||
{ apply eq_of_iso, apply terminal_iso_terminal},
|
{ apply eq_of_iso, apply terminal_iso_terminal},
|
||||||
induction p, exact ap _ !is_hprop.elim
|
induction p, exact ap _ !is_prop.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
variable (D)
|
variable (D)
|
||||||
|
@ -79,12 +79,12 @@ namespace category
|
||||||
|
|
||||||
section
|
section
|
||||||
open pi
|
open pi
|
||||||
theorem is_hprop_has_limits_of_shape [instance] (D : Category) (I : Precategory)
|
theorem is_prop_has_limits_of_shape [instance] (D : Category) (I : Precategory)
|
||||||
: is_hprop (has_limits_of_shape D I) :=
|
: is_prop (has_limits_of_shape D I) :=
|
||||||
by apply is_trunc_pi; intro F; exact is_hprop_has_terminal_object (Category_cone F)
|
by apply is_trunc_pi; intro F; exact is_prop_has_terminal_object (Category_cone F)
|
||||||
|
|
||||||
local attribute is_complete [reducible]
|
local attribute is_complete [reducible]
|
||||||
theorem is_hprop_is_complete [instance] (D : Category) : is_hprop (is_complete D) := _
|
theorem is_prop_is_complete [instance] (D : Category) : is_prop (is_complete D) := _
|
||||||
end
|
end
|
||||||
|
|
||||||
variables {D I}
|
variables {D I}
|
||||||
|
|
|
@ -49,7 +49,7 @@ namespace category
|
||||||
apply eq_of_homotopy, intro x, fapply sigma_eq: esimp,
|
apply eq_of_homotopy, intro x, fapply sigma_eq: esimp,
|
||||||
{ apply eq_of_homotopy, intro i, exact (ap10 (q i) x)⁻¹},
|
{ apply eq_of_homotopy, intro i, exact (ap10 (q i) x)⁻¹},
|
||||||
{ with_options [elaborator.ignore_instances true] -- TODO: fix
|
{ with_options [elaborator.ignore_instances true] -- TODO: fix
|
||||||
( refine is_hprop.elimo _ _ _;
|
( refine is_prop.elimo _ _ _;
|
||||||
refine is_trunc_pi _ _; intro i;
|
refine is_trunc_pi _ _; intro i;
|
||||||
refine is_trunc_pi _ _; intro j;
|
refine is_trunc_pi _ _; intro j;
|
||||||
refine is_trunc_pi _ _; intro f;
|
refine is_trunc_pi _ _; intro f;
|
||||||
|
@ -73,7 +73,7 @@ namespace category
|
||||||
fapply cone_obj.mk,
|
fapply cone_obj.mk,
|
||||||
{ fapply trunctype.mk,
|
{ fapply trunctype.mk,
|
||||||
{ apply set_quotient (is_cocomplete_set_cone_rel.{u v w} I F)},
|
{ apply set_quotient (is_cocomplete_set_cone_rel.{u v w} I F)},
|
||||||
{ apply is_hset_set_quotient}},
|
{ apply is_set_set_quotient}},
|
||||||
{ fapply nat_trans.mk,
|
{ fapply nat_trans.mk,
|
||||||
{ intro i x, esimp, apply class_of, exact ⟨i, x⟩},
|
{ intro i x, esimp, apply class_of, exact ⟨i, x⟩},
|
||||||
{ intro i j f, esimp, apply eq_of_homotopy, intro y, apply eq_of_rel, esimp,
|
{ intro i j f, esimp, apply eq_of_homotopy, intro y, apply eq_of_rel, esimp,
|
||||||
|
@ -99,7 +99,7 @@ namespace category
|
||||||
{ esimp, intro h, induction h with f q, apply cone_hom_eq, esimp at *,
|
{ esimp, intro h, induction h with f q, apply cone_hom_eq, esimp at *,
|
||||||
apply eq_of_homotopy, refine set_quotient.rec _ _,
|
apply eq_of_homotopy, refine set_quotient.rec _ _,
|
||||||
{ intro v, induction v with i x, esimp, exact (ap10 (q i) x)⁻¹},
|
{ intro v, induction v with i x, esimp, exact (ap10 (q i) x)⁻¹},
|
||||||
{ intro v w r, apply is_hprop.elimo}}},
|
{ intro v w r, apply is_prop.elimo}}},
|
||||||
end
|
end
|
||||||
|
|
||||||
end category
|
end category
|
||||||
|
|
|
@ -52,7 +52,7 @@ namespace nat_trans
|
||||||
(nat₂ : Π (a b : C) (f : hom a b), G f ∘ η₂ a = η₂ b ∘ F f)
|
(nat₂ : Π (a b : C) (f : hom a b), G f ∘ η₂ a = η₂ b ∘ F f)
|
||||||
(p : η₁ ~ η₂)
|
(p : η₁ ~ η₂)
|
||||||
: nat_trans.mk η₁ nat₁ = nat_trans.mk η₂ nat₂ :=
|
: nat_trans.mk η₁ nat₁ = nat_trans.mk η₂ nat₂ :=
|
||||||
apd011 nat_trans.mk (eq_of_homotopy p) !is_hprop.elim
|
apd011 nat_trans.mk (eq_of_homotopy p) !is_prop.elim
|
||||||
|
|
||||||
definition nat_trans_eq {η₁ η₂ : F ⟹ G} : natural_map η₁ ~ natural_map η₂ → η₁ = η₂ :=
|
definition nat_trans_eq {η₁ η₂ : F ⟹ G} : natural_map η₁ ~ natural_map η₂ → η₁ = η₂ :=
|
||||||
by induction η₁; induction η₂; apply nat_trans_mk_eq
|
by induction η₁; induction η₂; apply nat_trans_mk_eq
|
||||||
|
@ -82,10 +82,10 @@ namespace nat_trans
|
||||||
intro S,
|
intro S,
|
||||||
fapply sigma_eq,
|
fapply sigma_eq,
|
||||||
{ apply eq_of_homotopy, intro a, apply idp},
|
{ apply eq_of_homotopy, intro a, apply idp},
|
||||||
{ apply is_hprop.elimo}
|
{ apply is_prop.elimo}
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_hset_nat_trans [instance] : is_hset (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)
|
by apply is_trunc_is_equiv_closed; apply (equiv.to_is_equiv !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)
|
||||||
|
|
|
@ -28,10 +28,10 @@ namespace category
|
||||||
(id_left : Π ⦃a b : ob⦄ (f : hom a b), comp !ID f = f)
|
(id_left : Π ⦃a b : ob⦄ (f : hom a b), comp !ID f = f)
|
||||||
(id_right : Π ⦃a b : ob⦄ (f : hom a b), comp f !ID = f)
|
(id_right : Π ⦃a b : ob⦄ (f : hom a b), comp f !ID = f)
|
||||||
(id_id : Π (a : ob), comp !ID !ID = ID a)
|
(id_id : Π (a : ob), comp !ID !ID = ID a)
|
||||||
(is_hset_hom : Π(a b : ob), is_hset (hom a b))
|
(is_set_hom : Π(a b : ob), is_set (hom a b))
|
||||||
|
|
||||||
-- attribute precategory [multiple-instances] --this is not used anywhere
|
-- attribute precategory [multiple-instances] --this is not used anywhere
|
||||||
attribute precategory.is_hset_hom [instance]
|
attribute precategory.is_set_hom [instance]
|
||||||
|
|
||||||
infixr ∘ := precategory.comp
|
infixr ∘ := precategory.comp
|
||||||
-- input ⟶ using \--> (this is a different arrow than \-> (→))
|
-- input ⟶ using \--> (this is a different arrow than \-> (→))
|
||||||
|
@ -48,15 +48,15 @@ namespace category
|
||||||
abbreviation id_left [unfold 2] := @precategory.id_left
|
abbreviation id_left [unfold 2] := @precategory.id_left
|
||||||
abbreviation id_right [unfold 2] := @precategory.id_right
|
abbreviation id_right [unfold 2] := @precategory.id_right
|
||||||
abbreviation id_id [unfold 2] := @precategory.id_id
|
abbreviation id_id [unfold 2] := @precategory.id_id
|
||||||
abbreviation is_hset_hom [unfold 2] := @precategory.is_hset_hom
|
abbreviation is_set_hom [unfold 2] := @precategory.is_set_hom
|
||||||
|
|
||||||
definition is_hprop_hom_eq {ob : Type} [C : precategory ob] {x y : ob} (f g : x ⟶ y)
|
definition is_prop_hom_eq {ob : Type} [C : precategory ob] {x y : ob} (f g : x ⟶ y)
|
||||||
: is_hprop (f = g) :=
|
: is_prop (f = g) :=
|
||||||
_
|
_
|
||||||
|
|
||||||
-- the constructor you want to use in practice
|
-- the constructor you want to use in practice
|
||||||
protected definition precategory.mk [constructor] {ob : Type} (hom : ob → ob → Type)
|
protected definition precategory.mk [constructor] {ob : Type} (hom : ob → ob → Type)
|
||||||
[hset : Π (a b : ob), is_hset (hom a b)]
|
[hset : Π (a b : ob), is_set (hom a b)]
|
||||||
(comp : Π ⦃a b c : ob⦄, hom b c → hom a b → hom a c) (ID : Π (a : ob), hom a a)
|
(comp : Π ⦃a b c : ob⦄, hom b c → hom a b → hom a c) (ID : Π (a : ob), hom a a)
|
||||||
(ass : Π ⦃a b c d : ob⦄ (h : hom c d) (g : hom b c) (f : hom a b),
|
(ass : Π ⦃a b c d : ob⦄ (h : hom c d) (g : hom b c) (f : hom a b),
|
||||||
comp h (comp g f) = comp (comp h g) f)
|
comp h (comp g f) = comp (comp h g) f)
|
||||||
|
@ -179,7 +179,7 @@ namespace category
|
||||||
assert K : ID1 = ID2,
|
assert K : ID1 = ID2,
|
||||||
{ apply eq_of_homotopy, intro a, exact !ir'⁻¹ ⬝ !il},
|
{ apply eq_of_homotopy, intro a, exact !ir'⁻¹ ⬝ !il},
|
||||||
induction K,
|
induction K,
|
||||||
apply ap0111111 (precategory.mk' hom1 comp1 ID1): apply is_hprop.elim
|
apply ap0111111 (precategory.mk' hom1 comp1 ID1): apply is_prop.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
||||||
|
@ -211,7 +211,7 @@ namespace category
|
||||||
assert K : ID1 = ID2,
|
assert K : ID1 = ID2,
|
||||||
{ apply eq_of_homotopy, intros, apply r},
|
{ apply eq_of_homotopy, intros, apply r},
|
||||||
induction H, induction K,
|
induction H, induction K,
|
||||||
apply ap0111111 (precategory.mk' hom1 comp1 ID1): apply is_hprop.elim
|
apply ap0111111 (precategory.mk' hom1 comp1 ID1): apply is_prop.elim
|
||||||
end
|
end
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
|
|
@ -10,12 +10,12 @@ open is_trunc eq
|
||||||
|
|
||||||
namespace category
|
namespace category
|
||||||
structure strict_precategory [class] (ob : Type) extends precategory ob :=
|
structure strict_precategory [class] (ob : Type) extends precategory ob :=
|
||||||
mk' :: (is_hset_ob : is_hset ob)
|
mk' :: (is_set_ob : is_set ob)
|
||||||
|
|
||||||
attribute strict_precategory.is_hset_ob [instance]
|
attribute strict_precategory.is_set_ob [instance]
|
||||||
|
|
||||||
definition strict_precategory.mk [reducible] {ob : Type} (C : precategory ob)
|
definition strict_precategory.mk [reducible] {ob : Type} (C : precategory ob)
|
||||||
(H : is_hset ob) : strict_precategory ob :=
|
(H : is_set ob) : strict_precategory ob :=
|
||||||
precategory.rec_on C strict_precategory.mk' H
|
precategory.rec_on C strict_precategory.mk' H
|
||||||
|
|
||||||
structure Strict_precategory : Type :=
|
structure Strict_precategory : Type :=
|
||||||
|
|
|
@ -19,10 +19,10 @@ variable {A : Type}
|
||||||
namespace algebra
|
namespace algebra
|
||||||
|
|
||||||
structure semigroup [class] (A : Type) extends has_mul A :=
|
structure semigroup [class] (A : Type) extends has_mul A :=
|
||||||
(is_hset_carrier : is_hset A)
|
(is_set_carrier : is_set A)
|
||||||
(mul_assoc : Πa b c, mul (mul a b) c = mul a (mul b c))
|
(mul_assoc : Πa b c, mul (mul a b) c = mul a (mul b c))
|
||||||
|
|
||||||
attribute semigroup.is_hset_carrier [instance]
|
attribute semigroup.is_set_carrier [instance]
|
||||||
|
|
||||||
definition mul.assoc [s : semigroup A] (a b c : A) : a * b * c = a * (b * c) :=
|
definition mul.assoc [s : semigroup A] (a b c : A) : a * b * c = a * (b * c) :=
|
||||||
!semigroup.mul_assoc
|
!semigroup.mul_assoc
|
||||||
|
@ -60,10 +60,10 @@ abbreviation eq_of_mul_eq_mul_right' := @mul.right_cancel
|
||||||
/- additive semigroup -/
|
/- additive semigroup -/
|
||||||
|
|
||||||
structure add_semigroup [class] (A : Type) extends has_add A :=
|
structure add_semigroup [class] (A : Type) extends has_add A :=
|
||||||
(is_hset_carrier : is_hset A)
|
(is_set_carrier : is_set A)
|
||||||
(add_assoc : Πa b c, add (add a b) c = add a (add b c))
|
(add_assoc : Πa b c, add (add a b) c = add a (add b c))
|
||||||
|
|
||||||
attribute add_semigroup.is_hset_carrier [instance]
|
attribute add_semigroup.is_set_carrier [instance]
|
||||||
|
|
||||||
definition add.assoc [s : add_semigroup A] (a b c : A) : a + b + c = a + (b + c) :=
|
definition add.assoc [s : add_semigroup A] (a b c : A) : a + b + c = a + (b + c) :=
|
||||||
!add_semigroup.add_assoc
|
!add_semigroup.add_assoc
|
||||||
|
@ -128,7 +128,7 @@ definition add_monoid.to_monoid {A : Type} [s : add_monoid A] : monoid A :=
|
||||||
one := add_monoid.zero A,
|
one := add_monoid.zero A,
|
||||||
mul_one := add_monoid.add_zero,
|
mul_one := add_monoid.add_zero,
|
||||||
one_mul := add_monoid.zero_add,
|
one_mul := add_monoid.zero_add,
|
||||||
is_hset_carrier := _
|
is_set_carrier := _
|
||||||
⦄
|
⦄
|
||||||
|
|
||||||
definition add_comm_monoid.to_comm_monoid {A : Type} [s : add_comm_monoid A] : comm_monoid A :=
|
definition add_comm_monoid.to_comm_monoid {A : Type} [s : add_comm_monoid A] : comm_monoid A :=
|
||||||
|
@ -585,7 +585,7 @@ definition group_of_add_group (A : Type) [G : add_group A] : group A :=
|
||||||
mul_one := add_zero,
|
mul_one := add_zero,
|
||||||
inv := has_neg.neg,
|
inv := has_neg.neg,
|
||||||
mul_left_inv := add.left_inv,
|
mul_left_inv := add.left_inv,
|
||||||
is_hset_carrier := _⦄
|
is_set_carrier := _⦄
|
||||||
|
|
||||||
namespace norm_num
|
namespace norm_num
|
||||||
reveal add.assoc
|
reveal add.assoc
|
||||||
|
|
|
@ -50,12 +50,12 @@ namespace eq
|
||||||
prefix `π₁`:95 := fundamental_group
|
prefix `π₁`:95 := fundamental_group
|
||||||
|
|
||||||
open equiv unit
|
open equiv unit
|
||||||
theorem trivial_homotopy_of_is_hset (A : Type*) [H : is_hset A] (n : ℕ) : πG[n+1] A = G0 :=
|
theorem trivial_homotopy_of_is_set (A : Type*) [H : is_set A] (n : ℕ) : πG[n+1] A = G0 :=
|
||||||
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,
|
||||||
apply is_trunc_succ_succ_of_is_hset
|
apply is_trunc_succ_succ_of_is_set
|
||||||
end
|
end
|
||||||
|
|
||||||
definition homotopy_group_succ_out (A : Type*) (n : ℕ) : πG[ n +1] A = π₁ Ω[n] A := idp
|
definition homotopy_group_succ_out (A : Type*) (n : ℕ) : πG[ n +1] A = π₁ Ω[n] A := idp
|
||||||
|
@ -77,8 +77,8 @@ namespace eq
|
||||||
exact ap (Group_homotopy_group n) !loop_space_succ_eq_in⁻¹}
|
exact ap (Group_homotopy_group n) !loop_space_succ_eq_in⁻¹}
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem trivial_homotopy_of_is_hset_loop_space {A : Type*} {n : ℕ} (m : ℕ) (H : is_hset (Ω[n] A))
|
theorem trivial_homotopy_of_is_set_loop_space {A : Type*} {n : ℕ} (m : ℕ) (H : is_set (Ω[n] A))
|
||||||
: πG[m+n+1] A = G0 :=
|
: πG[m+n+1] A = G0 :=
|
||||||
!homotopy_group_add ⬝ !trivial_homotopy_of_is_hset
|
!homotopy_group_add ⬝ !trivial_homotopy_of_is_set
|
||||||
|
|
||||||
end eq
|
end eq
|
||||||
|
|
|
@ -50,11 +50,11 @@ namespace algebra
|
||||||
... = Hm G1 (Hi g) : by rewrite Gh4
|
... = Hm G1 (Hi g) : by rewrite Gh4
|
||||||
... = Hi g : Gh2),
|
... = Hi g : Gh2),
|
||||||
cases same_one, cases same_inv,
|
cases same_one, cases same_inv,
|
||||||
have ps : Gs = Hs, from !is_hprop.elim,
|
have ps : Gs = Hs, from !is_prop.elim,
|
||||||
have ph1 : Gh1 = Hh1, from !is_hprop.elim,
|
have ph1 : Gh1 = Hh1, from !is_prop.elim,
|
||||||
have ph2 : Gh2 = Hh2, from !is_hprop.elim,
|
have ph2 : Gh2 = Hh2, from !is_prop.elim,
|
||||||
have ph3 : Gh3 = Hh3, from !is_hprop.elim,
|
have ph3 : Gh3 = Hh3, from !is_prop.elim,
|
||||||
have ph4 : Gh4 = Hh4, from !is_hprop.elim,
|
have ph4 : Gh4 = Hh4, from !is_prop.elim,
|
||||||
cases ps, cases ph1, cases ph2, cases ph3, cases ph4, reflexivity
|
cases ps, cases ph1, cases ph2, cases ph3, cases ph4, reflexivity
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -86,7 +86,7 @@ namespace algebra
|
||||||
mul_one := trunc_mul_one,
|
mul_one := trunc_mul_one,
|
||||||
inv := trunc_inv,
|
inv := trunc_inv,
|
||||||
mul_left_inv := trunc_mul_left_inv,
|
mul_left_inv := trunc_mul_left_inv,
|
||||||
is_hset_carrier := _⦄
|
is_set_carrier := _⦄
|
||||||
|
|
||||||
|
|
||||||
definition trunc_comm_group [constructor] (mul_comm : ∀a b, mul a b = mul b a) : comm_group G :=
|
definition trunc_comm_group [constructor] (mul_comm : ∀a b, mul a b = mul b a) : comm_group G :=
|
||||||
|
|
|
@ -7,22 +7,22 @@ namespace choice
|
||||||
|
|
||||||
-- 3.8.1. The axiom of choice.
|
-- 3.8.1. The axiom of choice.
|
||||||
definition AC [reducible] := Π (X : Type.{u}) (A : X -> Type.{u}) (P : Π x, A x -> Type.{u}),
|
definition AC [reducible] := Π (X : Type.{u}) (A : X -> Type.{u}) (P : Π x, A x -> Type.{u}),
|
||||||
is_hset X -> (Π x, is_hset (A x)) -> (Π x a, is_hprop (P x a)) ->
|
is_set X -> (Π x, is_set (A x)) -> (Π x a, is_prop (P x a)) ->
|
||||||
(Π x, ∥ Σ a, P x a ∥) -> ∥ Σ f, Π x, P x (f x) ∥
|
(Π x, ∥ Σ a, P x a ∥) -> ∥ Σ f, Π x, P x (f x) ∥
|
||||||
|
|
||||||
-- 3.8.3. Corresponds to the assertion that
|
-- 3.8.3. Corresponds to the assertion that
|
||||||
-- "the cartesian product of a family of nonempty sets is nonempty".
|
-- "the cartesian product of a family of nonempty sets is nonempty".
|
||||||
definition AC_cart [reducible] := Π (X : Type.{u}) (A : X -> Type.{u}),
|
definition AC_cart [reducible] := Π (X : Type.{u}) (A : X -> Type.{u}),
|
||||||
is_hset X -> (Π x, is_hset (A x)) -> (Π x, ∥ A x ∥) -> ∥ Π x, A x ∥
|
is_set X -> (Π x, is_set (A x)) -> (Π x, ∥ A x ∥) -> ∥ Π x, A x ∥
|
||||||
|
|
||||||
-- A slight variant of AC with a modified (equivalent) codomain.
|
-- A slight variant of AC with a modified (equivalent) codomain.
|
||||||
definition AC' [reducible] := Π (X : Type.{u}) (A : X -> Type.{u}) (P : Π x, A x -> Type.{u}),
|
definition AC' [reducible] := Π (X : Type.{u}) (A : X -> Type.{u}) (P : Π x, A x -> Type.{u}),
|
||||||
is_hset X -> (Π x, is_hset (A x)) -> (Π x a, is_hprop (P x a))
|
is_set X -> (Π x, is_set (A x)) -> (Π x a, is_prop (P x a))
|
||||||
-> (Π x, ∥ Σ a, P x a ∥) -> ∥ Π x, Σ a : A x, P x a ∥
|
-> (Π x, ∥ Σ a, P x a ∥) -> ∥ Π x, Σ a : A x, P x a ∥
|
||||||
|
|
||||||
-- The equivalence of AC and AC' follows from the equivalence of their codomains.
|
-- The equivalence of AC and AC' follows from the equivalence of their codomains.
|
||||||
definition AC_equiv_AC' : AC.{u} ≃ AC'.{u} :=
|
definition AC_equiv_AC' : AC.{u} ≃ AC'.{u} :=
|
||||||
equiv_of_is_hprop
|
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))
|
||||||
|
|
||||||
|
@ -39,7 +39,7 @@ namespace choice
|
||||||
|
|
||||||
-- Which is enough to show AC' ≃ AC_cart, since both are hprops.
|
-- Which is enough to show AC' ≃ AC_cart, since both are hprops.
|
||||||
definition AC'_equiv_AC_cart : AC'.{u} ≃ AC_cart.{u} :=
|
definition AC'_equiv_AC_cart : AC'.{u} ≃ AC_cart.{u} :=
|
||||||
equiv_of_is_hprop 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} :=
|
||||||
|
@ -52,21 +52,21 @@ namespace choice
|
||||||
|
|
||||||
definition Y : X -> Type.{1} := λ x, x0 = x
|
definition Y : X -> Type.{1} := λ x, x0 = x
|
||||||
|
|
||||||
definition not_is_hset_X : ¬ is_hset X :=
|
definition not_is_set_X : ¬ is_set X :=
|
||||||
begin
|
begin
|
||||||
intro H, apply not_is_hprop_bool_eq_bool,
|
intro H, apply not_is_prop_bool_eq_bool,
|
||||||
apply @is_trunc_equiv_closed (x0 = x0),
|
apply @is_trunc_equiv_closed (x0 = x0),
|
||||||
apply equiv.symm !equiv_subtype
|
apply equiv.symm !equiv_subtype
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_hset_x1 (x : X) : is_hset x.1 :=
|
definition is_set_x1 (x : X) : is_set x.1 :=
|
||||||
by cases x; induction a_1; cases a_1; exact _
|
by cases x; induction a_1; cases a_1; exact _
|
||||||
|
|
||||||
definition is_hset_Yx (x : X) : is_hset (Y x) :=
|
definition is_set_Yx (x : X) : is_set (Y x) :=
|
||||||
begin
|
begin
|
||||||
apply @is_trunc_equiv_closed _ _ _ !equiv_subtype,
|
apply @is_trunc_equiv_closed _ _ _ !equiv_subtype,
|
||||||
apply @is_trunc_equiv_closed _ _ _ (equiv.symm !eq_equiv_equiv),
|
apply @is_trunc_equiv_closed _ _ _ (equiv.symm !eq_equiv_equiv),
|
||||||
apply is_trunc_equiv; repeat (apply is_hset_x1)
|
apply is_trunc_equiv; repeat (apply is_set_x1)
|
||||||
end
|
end
|
||||||
|
|
||||||
definition trunc_Yx (x : X) : ∥ Y x ∥ :=
|
definition trunc_Yx (x : X) : ∥ Y x ∥ :=
|
||||||
|
@ -81,8 +81,8 @@ namespace choice
|
||||||
-- 3.8.5. There exists a type X and a family Y : X → U such that each Y(x) is a set,
|
-- 3.8.5. There exists a type X and a family Y : X → U such that each Y(x) is a set,
|
||||||
-- but such that (3.8.3) is false.
|
-- but such that (3.8.3) is false.
|
||||||
definition X_must_be_hset : Σ (X : Type.{1}) (Y : X -> Type.{1})
|
definition X_must_be_hset : Σ (X : Type.{1}) (Y : X -> Type.{1})
|
||||||
(HA : Π x : X, is_hset (Y x)), ¬ ((Π x : X, ∥ Y x ∥) -> ∥ Π x : X, Y x ∥) :=
|
(HA : Π x : X, is_set (Y x)), ¬ ((Π x : X, ∥ Y x ∥) -> ∥ Π x : X, Y x ∥) :=
|
||||||
⟨X, Y, is_hset_Yx, λ H, trunc.rec_on (H trunc_Yx)
|
⟨X, Y, is_set_Yx, λ H, trunc.rec_on (H trunc_Yx)
|
||||||
(λ H0, not_is_hset_X (@is_trunc_of_is_contr _ _ (is_contr.mk x0 H0)))⟩
|
(λ H0, not_is_set_X (@is_trunc_of_is_contr _ _ (is_contr.mk x0 H0)))⟩
|
||||||
|
|
||||||
end choice
|
end choice
|
||||||
|
|
|
@ -483,8 +483,8 @@ namespace eq
|
||||||
by induction s₂;induction s₁;constructor
|
by induction s₂;induction s₁;constructor
|
||||||
|
|
||||||
open is_trunc
|
open is_trunc
|
||||||
definition is_hset.elims [H : is_hset A] : square p₁₀ p₁₂ p₀₁ p₂₁ :=
|
definition is_set.elims [H : is_set A] : square p₁₀ p₁₂ p₀₁ p₂₁ :=
|
||||||
square_of_eq !is_hset.elim
|
square_of_eq !is_set.elim
|
||||||
|
|
||||||
-- definition square_of_con_inv_hsquare {p₁ p₂ p₃ p₄ : a₁ = a₂}
|
-- definition square_of_con_inv_hsquare {p₁ p₂ p₃ p₄ : a₁ = a₂}
|
||||||
-- {t : p₁ = p₂} {b : p₃ = p₄} {l : p₁ = p₃} {r : p₂ = p₄}
|
-- {t : p₁ = p₂} {b : p₃ = p₄} {l : p₁ = p₃} {r : p₂ = p₄}
|
||||||
|
|
|
@ -57,54 +57,54 @@ namespace function
|
||||||
: f a = f a' → a = a' :=
|
: f a = f a' → a = a' :=
|
||||||
(ap f)⁻¹
|
(ap f)⁻¹
|
||||||
|
|
||||||
definition is_embedding_of_is_injective [HA : is_hset A] [HB : is_hset B]
|
definition is_embedding_of_is_injective [HA : is_set A] [HB : is_set B]
|
||||||
(H : Π(a a' : A), f a = f a' → a = a') : is_embedding f :=
|
(H : Π(a a' : A), f a = f a' → a = a') : is_embedding f :=
|
||||||
begin
|
begin
|
||||||
intro a a',
|
intro a a',
|
||||||
fapply adjointify,
|
fapply adjointify,
|
||||||
{exact (H a a')},
|
{exact (H a a')},
|
||||||
{intro p, apply is_hset.elim},
|
{intro p, apply is_set.elim},
|
||||||
{intro p, apply is_hset.elim}
|
{intro p, apply is_set.elim}
|
||||||
end
|
end
|
||||||
|
|
||||||
variable (f)
|
variable (f)
|
||||||
|
|
||||||
definition is_hprop_is_embedding [instance] : is_hprop (is_embedding f) :=
|
definition is_prop_is_embedding [instance] : is_prop (is_embedding f) :=
|
||||||
by unfold is_embedding; exact _
|
by unfold is_embedding; exact _
|
||||||
|
|
||||||
definition is_embedding_equiv_is_injective [HA : is_hset A] [HB : is_hset B]
|
definition is_embedding_equiv_is_injective [HA : is_set A] [HB : is_set B]
|
||||||
: is_embedding f ≃ (Π(a a' : A), f a = f a' → a = a') :=
|
: is_embedding f ≃ (Π(a a' : A), f a = f a' → a = a') :=
|
||||||
begin
|
begin
|
||||||
fapply equiv.MK,
|
fapply equiv.MK,
|
||||||
{ apply @is_injective_of_is_embedding},
|
{ apply @is_injective_of_is_embedding},
|
||||||
{ apply is_embedding_of_is_injective},
|
{ apply is_embedding_of_is_injective},
|
||||||
{ intro H, apply is_hprop.elim},
|
{ intro H, apply is_prop.elim},
|
||||||
{ intro H, apply is_hprop.elim, }
|
{ intro H, apply is_prop.elim, }
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_hprop_fiber_of_is_embedding [H : is_embedding f] (b : B) :
|
definition is_prop_fiber_of_is_embedding [H : is_embedding f] (b : B) :
|
||||||
is_hprop (fiber f b) :=
|
is_prop (fiber f b) :=
|
||||||
begin
|
begin
|
||||||
apply is_hprop.mk, intro v w,
|
apply is_prop.mk, intro v w,
|
||||||
induction v with a p, induction w with a' q, induction q,
|
induction v with a p, induction w with a' q, induction q,
|
||||||
fapply fiber_eq,
|
fapply fiber_eq,
|
||||||
{ esimp, apply is_injective_of_is_embedding p},
|
{ esimp, apply is_injective_of_is_embedding p},
|
||||||
{ esimp [is_injective_of_is_embedding], symmetry, apply right_inv}
|
{ esimp [is_injective_of_is_embedding], symmetry, apply right_inv}
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_hprop_fun_of_is_embedding [H : is_embedding f] : is_trunc_fun -1 f :=
|
definition is_prop_fun_of_is_embedding [H : is_embedding f] : is_trunc_fun -1 f :=
|
||||||
is_hprop_fiber_of_is_embedding f
|
is_prop_fiber_of_is_embedding f
|
||||||
|
|
||||||
definition is_embedding_of_is_hprop_fun [constructor] [H : is_trunc_fun -1 f] : is_embedding f :=
|
definition is_embedding_of_is_prop_fun [constructor] [H : is_trunc_fun -1 f] : is_embedding f :=
|
||||||
begin
|
begin
|
||||||
intro a a', fapply adjointify,
|
intro a a', fapply adjointify,
|
||||||
{ intro p, exact ap point (@is_hprop.elim (fiber f (f a')) _ (fiber.mk a p) (fiber.mk a' idp))},
|
{ intro p, exact ap point (@is_prop.elim (fiber f (f a')) _ (fiber.mk a p) (fiber.mk a' idp))},
|
||||||
{ intro p, rewrite [-ap_compose], esimp, apply ap_con_eq (@point_eq _ _ f (f a'))},
|
{ intro p, rewrite [-ap_compose], esimp, apply ap_con_eq (@point_eq _ _ f (f a'))},
|
||||||
{ intro p, induction p, apply ap (ap point), apply is_hprop_elim_self}
|
{ intro p, induction p, apply ap (ap point), apply is_prop_elim_self}
|
||||||
end
|
end
|
||||||
|
|
||||||
variable {f}
|
variable {f}
|
||||||
definition is_surjective_rec_on {P : Type} (H : is_surjective f) (b : B) [Pt : is_hprop P]
|
definition is_surjective_rec_on {P : Type} (H : is_surjective f) (b : B) [Pt : is_prop P]
|
||||||
(IH : fiber f b → P) : P :=
|
(IH : fiber f b → P) : P :=
|
||||||
trunc.rec_on (H b) IH
|
trunc.rec_on (H b) IH
|
||||||
variable (f)
|
variable (f)
|
||||||
|
@ -113,7 +113,7 @@ namespace function
|
||||||
: is_surjective f :=
|
: is_surjective f :=
|
||||||
λb, tr (H b)
|
λb, tr (H b)
|
||||||
|
|
||||||
definition is_hprop_is_surjective [instance] : is_hprop (is_surjective f) :=
|
definition is_prop_is_surjective [instance] : is_prop (is_surjective f) :=
|
||||||
by unfold is_surjective; exact _
|
by unfold is_surjective; exact _
|
||||||
|
|
||||||
definition is_weakly_constant_ap [instance] [H : is_weakly_constant f] (a a' : A) :
|
definition is_weakly_constant_ap [instance] [H : is_weakly_constant f] (a a' : A) :
|
||||||
|
@ -184,8 +184,8 @@ namespace function
|
||||||
: is_constant (f ∘ point : fiber f b → B) :=
|
: is_constant (f ∘ point : fiber f b → B) :=
|
||||||
is_constant.mk b (λv, by induction v with a p;exact p)
|
is_constant.mk b (λv, by induction v with a p;exact p)
|
||||||
|
|
||||||
definition is_embedding_of_is_hprop_fiber [H : Π(b : B), is_hprop (fiber f b)] : is_embedding f :=
|
definition is_embedding_of_is_prop_fiber [H : Π(b : B), is_prop (fiber f b)] : is_embedding f :=
|
||||||
is_embedding_of_is_hprop_fun _
|
is_embedding_of_is_prop_fun _
|
||||||
|
|
||||||
definition is_retraction_of_is_equiv [instance] [H : is_equiv f] : is_retraction f :=
|
definition is_retraction_of_is_equiv [instance] [H : is_equiv f] : is_retraction f :=
|
||||||
is_retraction.mk f⁻¹ (right_inv f)
|
is_retraction.mk f⁻¹ (right_inv f)
|
||||||
|
@ -208,9 +208,9 @@ namespace function
|
||||||
local attribute is_equiv_of_is_section_of_is_retraction [instance] [priority 10000]
|
local attribute is_equiv_of_is_section_of_is_retraction [instance] [priority 10000]
|
||||||
local attribute trunctype.struct [instance] [priority 1] -- remove after #842 is closed
|
local attribute trunctype.struct [instance] [priority 1] -- remove after #842 is closed
|
||||||
variable (f)
|
variable (f)
|
||||||
definition is_hprop_is_retraction_prod_is_section : is_hprop (is_retraction f × is_section f) :=
|
definition is_prop_is_retraction_prod_is_section : is_prop (is_retraction f × is_section f) :=
|
||||||
begin
|
begin
|
||||||
apply is_hprop_of_imp_is_contr, intro H, induction H with H1 H2,
|
apply is_prop_of_imp_is_contr, intro H, induction H with H1 H2,
|
||||||
exact _,
|
exact _,
|
||||||
end
|
end
|
||||||
end
|
end
|
||||||
|
@ -233,11 +233,11 @@ namespace function
|
||||||
exact ap r (center_eq (is_retraction.sect r b)) ⬝ (is_retraction.right_inverse r b)
|
exact ap r (center_eq (is_retraction.sect r b)) ⬝ (is_retraction.right_inverse r b)
|
||||||
end
|
end
|
||||||
|
|
||||||
local attribute is_hprop_is_retraction_prod_is_section [instance]
|
local attribute is_prop_is_retraction_prod_is_section [instance]
|
||||||
definition is_retraction_prod_is_section_equiv_is_equiv [constructor]
|
definition is_retraction_prod_is_section_equiv_is_equiv [constructor]
|
||||||
: (is_retraction f × is_section f) ≃ is_equiv f :=
|
: (is_retraction f × is_section f) ≃ is_equiv f :=
|
||||||
begin
|
begin
|
||||||
apply equiv_of_is_hprop,
|
apply equiv_of_is_prop,
|
||||||
intro H, induction H, apply is_equiv_of_is_section_of_is_retraction,
|
intro H, induction H, apply is_equiv_of_is_section_of_is_retraction,
|
||||||
intro H, split, repeat exact _
|
intro H, split, repeat exact _
|
||||||
end
|
end
|
||||||
|
|
|
@ -74,12 +74,12 @@ parameters {A B : Type.{u}} (f g : A → B)
|
||||||
(x : A) : transport (elim_type P_i Pcp) (cp x) = Pcp x :=
|
(x : A) : transport (elim_type P_i Pcp) (cp x) = Pcp x :=
|
||||||
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_cp];apply cast_ua_fn
|
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_cp];apply cast_ua_fn
|
||||||
|
|
||||||
protected definition rec_hprop {P : coeq → Type} [H : Πx, is_hprop (P x)]
|
protected definition rec_hprop {P : coeq → Type} [H : Πx, is_prop (P x)]
|
||||||
(P_i : Π(x : B), P (coeq_i x)) (y : coeq) : P y :=
|
(P_i : Π(x : B), P (coeq_i x)) (y : coeq) : P y :=
|
||||||
rec P_i (λa, !is_hprop.elimo) y
|
rec P_i (λa, !is_prop.elimo) y
|
||||||
|
|
||||||
protected definition elim_hprop {P : Type} [H : is_hprop P] (P_i : B → P) (y : coeq) : P :=
|
protected definition elim_hprop {P : Type} [H : is_prop P] (P_i : B → P) (y : coeq) : P :=
|
||||||
elim P_i (λa, !is_hprop.elim) y
|
elim P_i (λa, !is_prop.elim) y
|
||||||
|
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -85,13 +85,13 @@ section
|
||||||
{j : J} (x : A (dom j)) : transport (elim_type Pincl Pglue) (cglue j x) = Pglue j x :=
|
{j : J} (x : A (dom j)) : transport (elim_type Pincl Pglue) (cglue j x) = Pglue j x :=
|
||||||
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_cglue];apply cast_ua_fn
|
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_cglue];apply cast_ua_fn
|
||||||
|
|
||||||
protected definition rec_hprop {P : colimit → Type} [H : Πx, is_hprop (P x)]
|
protected definition rec_hprop {P : colimit → Type} [H : Πx, is_prop (P x)]
|
||||||
(Pincl : Π⦃i : I⦄ (x : A i), P (ι x)) (y : colimit) : P y :=
|
(Pincl : Π⦃i : I⦄ (x : A i), P (ι x)) (y : colimit) : P y :=
|
||||||
rec Pincl (λa b, !is_hprop.elimo) y
|
rec Pincl (λa b, !is_prop.elimo) y
|
||||||
|
|
||||||
protected definition elim_hprop {P : Type} [H : is_hprop P] (Pincl : Π⦃i : I⦄ (x : A i), P)
|
protected definition elim_hprop {P : Type} [H : is_prop P] (Pincl : Π⦃i : I⦄ (x : A i), P)
|
||||||
(y : colimit) : P :=
|
(y : colimit) : P :=
|
||||||
elim Pincl (λa b, !is_hprop.elim) y
|
elim Pincl (λa b, !is_prop.elim) y
|
||||||
|
|
||||||
end
|
end
|
||||||
end colimit
|
end colimit
|
||||||
|
@ -175,13 +175,13 @@ section
|
||||||
: transport (elim_type Pincl Pglue) (glue a) = Pglue a :=
|
: transport (elim_type Pincl Pglue) (glue a) = Pglue a :=
|
||||||
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_glue];apply cast_ua_fn
|
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_glue];apply cast_ua_fn
|
||||||
|
|
||||||
protected definition rec_hprop {P : seq_colim → Type} [H : Πx, is_hprop (P x)]
|
protected definition rec_hprop {P : seq_colim → Type} [H : Πx, is_prop (P x)]
|
||||||
(Pincl : Π⦃n : ℕ⦄ (a : A n), P (sι a)) (aa : seq_colim) : P aa :=
|
(Pincl : Π⦃n : ℕ⦄ (a : A n), P (sι a)) (aa : seq_colim) : P aa :=
|
||||||
rec Pincl (λa b, !is_hprop.elimo) aa
|
rec Pincl (λa b, !is_prop.elimo) aa
|
||||||
|
|
||||||
protected definition elim_hprop {P : Type} [H : is_hprop P] (Pincl : Π⦃n : ℕ⦄ (a : A n), P)
|
protected definition elim_hprop {P : Type} [H : is_prop P] (Pincl : Π⦃n : ℕ⦄ (a : A n), P)
|
||||||
: seq_colim → P :=
|
: seq_colim → P :=
|
||||||
elim Pincl (λa b, !is_hprop.elim)
|
elim Pincl (λa b, !is_prop.elim)
|
||||||
|
|
||||||
|
|
||||||
end
|
end
|
||||||
|
|
|
@ -83,13 +83,13 @@ parameters {TL BL TR : Type} (f : TL → BL) (g : TL → TR)
|
||||||
: transport (elim_type Pinl Pinr Pglue) (glue x) = Pglue x :=
|
: transport (elim_type Pinl Pinr Pglue) (glue x) = Pglue x :=
|
||||||
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_glue];apply cast_ua_fn
|
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_glue];apply cast_ua_fn
|
||||||
|
|
||||||
protected definition rec_hprop {P : pushout → Type} [H : Πx, is_hprop (P x)]
|
protected definition rec_hprop {P : pushout → Type} [H : Πx, is_prop (P x)]
|
||||||
(Pinl : Π(x : BL), P (inl x)) (Pinr : Π(x : TR), P (inr x)) (y : pushout) :=
|
(Pinl : Π(x : BL), P (inl x)) (Pinr : Π(x : TR), P (inr x)) (y : pushout) :=
|
||||||
rec Pinl Pinr (λx, !is_hprop.elimo) y
|
rec Pinl Pinr (λx, !is_prop.elimo) y
|
||||||
|
|
||||||
protected definition elim_hprop {P : Type} [H : is_hprop P] (Pinl : BL → P) (Pinr : TR → P)
|
protected definition elim_hprop {P : Type} [H : is_prop P] (Pinl : BL → P) (Pinr : TR → P)
|
||||||
(y : pushout) : P :=
|
(y : pushout) : P :=
|
||||||
elim Pinl Pinr (λa, !is_hprop.elim) y
|
elim Pinl Pinr (λa, !is_prop.elim) y
|
||||||
|
|
||||||
end
|
end
|
||||||
end pushout
|
end pushout
|
||||||
|
|
|
@ -39,11 +39,11 @@ namespace quotient
|
||||||
end
|
end
|
||||||
|
|
||||||
protected definition rec_hprop {A : Type} {R : A → A → Type} {P : quotient R → Type}
|
protected definition rec_hprop {A : Type} {R : A → A → Type} {P : quotient R → Type}
|
||||||
[H : Πx, is_hprop (P x)] (Pc : Π(a : A), P (class_of R a)) (x : quotient R) : P x :=
|
[H : Πx, is_prop (P x)] (Pc : Π(a : A), P (class_of R a)) (x : quotient R) : P x :=
|
||||||
quotient.rec Pc (λa a' H, !is_hprop.elimo) x
|
quotient.rec Pc (λa a' H, !is_prop.elimo) x
|
||||||
|
|
||||||
protected definition elim_hprop {P : Type} [H : is_hprop P] (Pc : A → P) (x : quotient R) : P :=
|
protected definition elim_hprop {P : Type} [H : is_prop P] (Pc : A → P) (x : quotient R) : P :=
|
||||||
quotient.elim Pc (λa a' H, !is_hprop.elim) x
|
quotient.elim Pc (λa a' H, !is_prop.elim) x
|
||||||
|
|
||||||
protected definition elim_type (Pc : A → Type)
|
protected definition elim_type (Pc : A → Type)
|
||||||
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a') : quotient R → Type :=
|
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a') : quotient R → Type :=
|
||||||
|
|
|
@ -23,10 +23,10 @@ section
|
||||||
definition eq_of_rel {a a' : A} (H : R a a') : class_of a = class_of a' :=
|
definition eq_of_rel {a a' : A} (H : R a a') : class_of a = class_of a' :=
|
||||||
ap tr (eq_of_rel _ H)
|
ap tr (eq_of_rel _ H)
|
||||||
|
|
||||||
theorem is_hset_set_quotient [instance] : is_hset set_quotient :=
|
theorem is_set_set_quotient [instance] : is_set set_quotient :=
|
||||||
begin unfold set_quotient, exact _ end
|
begin unfold set_quotient, exact _ end
|
||||||
|
|
||||||
protected definition rec {P : set_quotient → Type} [Pt : Πaa, is_hset (P aa)]
|
protected definition rec {P : set_quotient → Type} [Pt : Πaa, is_set (P aa)]
|
||||||
(Pc : Π(a : A), P (class_of a)) (Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a =[eq_of_rel H] Pc a')
|
(Pc : Π(a : A), P (class_of a)) (Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a =[eq_of_rel H] Pc a')
|
||||||
(x : set_quotient) : P x :=
|
(x : set_quotient) : P x :=
|
||||||
begin
|
begin
|
||||||
|
@ -38,24 +38,24 @@ section
|
||||||
end
|
end
|
||||||
|
|
||||||
protected definition rec_on [reducible] {P : set_quotient → Type} (x : set_quotient)
|
protected definition rec_on [reducible] {P : set_quotient → Type} (x : set_quotient)
|
||||||
[Pt : Πaa, is_hset (P aa)] (Pc : Π(a : A), P (class_of a))
|
[Pt : Πaa, is_set (P aa)] (Pc : Π(a : A), P (class_of a))
|
||||||
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a =[eq_of_rel H] Pc a') : P x :=
|
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a =[eq_of_rel H] Pc a') : P x :=
|
||||||
rec Pc Pp x
|
rec Pc Pp x
|
||||||
|
|
||||||
theorem rec_eq_of_rel {P : set_quotient → Type} [Pt : Πaa, is_hset (P aa)]
|
theorem rec_eq_of_rel {P : set_quotient → Type} [Pt : Πaa, is_set (P aa)]
|
||||||
(Pc : Π(a : A), P (class_of a)) (Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a =[eq_of_rel H] Pc a')
|
(Pc : Π(a : A), P (class_of a)) (Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a =[eq_of_rel H] Pc a')
|
||||||
{a a' : A} (H : R a a') : apdo (rec Pc Pp) (eq_of_rel H) = Pp H :=
|
{a a' : A} (H : R a a') : apdo (rec Pc Pp) (eq_of_rel H) = Pp H :=
|
||||||
!is_hset.elimo
|
!is_set.elimo
|
||||||
|
|
||||||
protected definition elim {P : Type} [Pt : is_hset P] (Pc : A → P)
|
protected definition elim {P : Type} [Pt : is_set P] (Pc : A → P)
|
||||||
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a') (x : set_quotient) : P :=
|
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a') (x : set_quotient) : P :=
|
||||||
rec Pc (λa a' H, pathover_of_eq (Pp H)) x
|
rec Pc (λa a' H, pathover_of_eq (Pp H)) x
|
||||||
|
|
||||||
protected definition elim_on [reducible] {P : Type} (x : set_quotient) [Pt : is_hset P]
|
protected definition elim_on [reducible] {P : Type} (x : set_quotient) [Pt : is_set P]
|
||||||
(Pc : A → P) (Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a') : P :=
|
(Pc : A → P) (Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a') : P :=
|
||||||
elim Pc Pp x
|
elim Pc Pp x
|
||||||
|
|
||||||
theorem elim_eq_of_rel {P : Type} [Pt : is_hset P] (Pc : A → P)
|
theorem elim_eq_of_rel {P : Type} [Pt : is_set P] (Pc : A → P)
|
||||||
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a') {a a' : A} (H : R a a')
|
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a') {a a' : A} (H : R a a')
|
||||||
: ap (elim Pc Pp) (eq_of_rel H) = Pp H :=
|
: ap (elim Pc Pp) (eq_of_rel H) = Pp H :=
|
||||||
begin
|
begin
|
||||||
|
@ -63,13 +63,13 @@ section
|
||||||
rewrite [▸*,-apdo_eq_pathover_of_eq_ap,↑elim,rec_eq_of_rel],
|
rewrite [▸*,-apdo_eq_pathover_of_eq_ap,↑elim,rec_eq_of_rel],
|
||||||
end
|
end
|
||||||
|
|
||||||
protected definition rec_hprop {P : set_quotient → Type} [Pt : Πaa, is_hprop (P aa)]
|
protected definition rec_hprop {P : set_quotient → Type} [Pt : Πaa, is_prop (P aa)]
|
||||||
(Pc : Π(a : A), P (class_of a)) (x : set_quotient) : P x :=
|
(Pc : Π(a : A), P (class_of a)) (x : set_quotient) : P x :=
|
||||||
rec Pc (λa a' H, !is_hprop.elimo) x
|
rec Pc (λa a' H, !is_prop.elimo) x
|
||||||
|
|
||||||
protected definition elim_hprop {P : Type} [Pt : is_hprop P] (Pc : A → P) (x : set_quotient)
|
protected definition elim_hprop {P : Type} [Pt : is_prop P] (Pc : A → P) (x : set_quotient)
|
||||||
: P :=
|
: P :=
|
||||||
elim Pc (λa a' H, !is_hprop.elim) x
|
elim Pc (λa a' H, !is_prop.elim) x
|
||||||
|
|
||||||
end
|
end
|
||||||
end set_quotient
|
end set_quotient
|
||||||
|
@ -84,11 +84,11 @@ namespace set_quotient
|
||||||
variables {A B C : Type} (R : A → A → hprop) (S : B → B → hprop) (T : C → C → hprop)
|
variables {A B C : Type} (R : A → A → hprop) (S : B → B → hprop) (T : C → C → hprop)
|
||||||
|
|
||||||
definition is_surjective_class_of : is_surjective (class_of : A → set_quotient R) :=
|
definition is_surjective_class_of : is_surjective (class_of : A → set_quotient R) :=
|
||||||
λx, set_quotient.rec_on x (λa, tr (fiber.mk a idp)) (λa a' r, !is_hprop.elimo)
|
λx, set_quotient.rec_on x (λa, tr (fiber.mk a idp)) (λa a' r, !is_prop.elimo)
|
||||||
|
|
||||||
/- non-dependent universal property -/
|
/- non-dependent universal property -/
|
||||||
|
|
||||||
definition set_quotient_arrow_equiv (B : Type) [H : is_hset B] :
|
definition set_quotient_arrow_equiv (B : Type) [H : is_set B] :
|
||||||
(set_quotient R → B) ≃ (Σ(f : A → B), Π(a a' : A), R a a' → f a = f a') :=
|
(set_quotient R → B) ≃ (Σ(f : A → B), Π(a a' : A), R a a' → f a = f a') :=
|
||||||
begin
|
begin
|
||||||
fapply equiv.MK,
|
fapply equiv.MK,
|
||||||
|
@ -108,7 +108,7 @@ namespace set_quotient
|
||||||
intros a' a'' H1,
|
intros a' a'' H1,
|
||||||
refine to_inv !trunctype_eq_equiv _, esimp,
|
refine to_inv !trunctype_eq_equiv _, esimp,
|
||||||
apply ua,
|
apply ua,
|
||||||
apply equiv_of_is_hprop,
|
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
|
||||||
|
@ -139,7 +139,7 @@ namespace set_quotient
|
||||||
{ intro b b' s, apply eq_of_rel, apply H, apply is_reflexive.refl R, exact s}},
|
{ intro b b' s, apply eq_of_rel, apply H, apply is_reflexive.refl R, exact s}},
|
||||||
{ intro a a' r, apply eq_of_homotopy, refine set_quotient.rec _ _,
|
{ intro a a' r, apply eq_of_homotopy, refine set_quotient.rec _ _,
|
||||||
{ intro b, esimp, apply eq_of_rel, apply H, exact r, apply is_reflexive.refl S},
|
{ intro b, esimp, apply eq_of_rel, apply H, exact r, apply is_reflexive.refl S},
|
||||||
{ intro b b' s, apply eq_pathover, esimp, apply is_hset.elims}}
|
{ intro b b' s, apply eq_pathover, esimp, apply is_set.elims}}
|
||||||
end
|
end
|
||||||
|
|
||||||
end set_quotient
|
end set_quotient
|
||||||
|
|
|
@ -117,7 +117,7 @@ namespace trunc
|
||||||
definition or.intro_left [reducible] [constructor] (x : X) : X ∨ Y := tr (inl x)
|
definition or.intro_left [reducible] [constructor] (x : X) : X ∨ Y := tr (inl x)
|
||||||
definition or.intro_right [reducible] [constructor] (y : Y) : X ∨ Y := tr (inr y)
|
definition or.intro_right [reducible] [constructor] (y : Y) : X ∨ Y := tr (inr y)
|
||||||
|
|
||||||
definition is_contr_of_merely_hprop [H : is_hprop A] (aa : merely A) : is_contr A :=
|
definition is_contr_of_merely_hprop [H : is_prop A] (aa : merely A) : is_contr A :=
|
||||||
is_contr_of_inhabited_hprop (trunc.rec_on aa id)
|
is_contr_of_inhabited_hprop (trunc.rec_on aa id)
|
||||||
|
|
||||||
section
|
section
|
||||||
|
|
|
@ -219,7 +219,7 @@ namespace circle
|
||||||
exact ap succ p},
|
exact ap succ p},
|
||||||
{ intros n p, rewrite [↑circle.encode, nat_succ_eq_int_succ, neg_succ, -power_con_inv,
|
{ intros n p, rewrite [↑circle.encode, nat_succ_eq_int_succ, neg_succ, -power_con_inv,
|
||||||
@con_tr _ circle.code, transport_code_loop_inv, ↑[circle.encode] at p, p, -neg_succ] }},
|
@con_tr _ circle.code, transport_code_loop_inv, ↑[circle.encode] at p, p, -neg_succ] }},
|
||||||
{ apply pathover_of_tr_eq, apply eq_of_homotopy, intro a, apply @is_hset.elim,
|
{ apply pathover_of_tr_eq, apply eq_of_homotopy, intro a, apply @is_set.elim,
|
||||||
esimp, exact _} end end},
|
esimp, exact _} end end},
|
||||||
{ intro p, cases p, exact idp},
|
{ intro p, cases p, exact idp},
|
||||||
end
|
end
|
||||||
|
@ -254,7 +254,7 @@ namespace circle
|
||||||
open nat
|
open nat
|
||||||
definition homotopy_group_of_circle (n : ℕ) : πG[n+1 +1] S¹. = G0 :=
|
definition homotopy_group_of_circle (n : ℕ) : πG[n+1 +1] S¹. = G0 :=
|
||||||
begin
|
begin
|
||||||
refine @trivial_homotopy_of_is_hset_loop_space S¹. 1 n _,
|
refine @trivial_homotopy_of_is_set_loop_space S¹. 1 n _,
|
||||||
apply is_trunc_equiv_closed_rev, apply base_eq_base_equiv
|
apply is_trunc_equiv_closed_rev, apply base_eq_base_equiv
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -255,7 +255,7 @@ namespace homotopy
|
||||||
apply pathover_of_tr_eq,
|
apply pathover_of_tr_eq,
|
||||||
rewrite [transport_eq_Fr,idp_con],
|
rewrite [transport_eq_Fr,idp_con],
|
||||||
revert H, induction n with [n, IH],
|
revert H, induction n with [n, IH],
|
||||||
{ intro H, apply is_hprop.elim },
|
{ intro H, apply is_prop.elim },
|
||||||
{ intros H,
|
{ intros H,
|
||||||
change ap (@tr n .+2 (susp A)) (merid a) = ap tr (merid a'),
|
change ap (@tr n .+2 (susp A)) (merid a) = ap tr (merid a'),
|
||||||
generalize a',
|
generalize a',
|
||||||
|
|
|
@ -190,7 +190,7 @@ namespace is_trunc
|
||||||
let H' := iff.elim_left (is_trunc_iff_is_contr_loop n A) H a,
|
let H' := iff.elim_left (is_trunc_iff_is_contr_loop n A) H a,
|
||||||
note H'' := @is_trunc_equiv_closed_rev _ _ _ !pmap_sphere H',
|
note H'' := @is_trunc_equiv_closed_rev _ _ _ !pmap_sphere H',
|
||||||
assert p : (f = pmap.mk (λx, f base) (respect_pt f)),
|
assert p : (f = pmap.mk (λx, f base) (respect_pt f)),
|
||||||
apply is_hprop.elim,
|
apply is_prop.elim,
|
||||||
exact ap10 (ap pmap.to_fun p) x
|
exact ap10 (ap pmap.to_fun p) x
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -35,16 +35,16 @@ namespace is_trunc
|
||||||
induction (P a b), apply idp},
|
induction (P a b), apply idp},
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_hprop_is_trunc [instance] (n : trunc_index) :
|
definition is_prop_is_trunc [instance] (n : trunc_index) :
|
||||||
Π (A : Type), is_hprop (is_trunc n A) :=
|
Π (A : Type), is_prop (is_trunc n A) :=
|
||||||
begin
|
begin
|
||||||
induction n,
|
induction n,
|
||||||
{ intro A,
|
{ intro A,
|
||||||
apply is_trunc_is_equiv_closed,
|
apply is_trunc_is_equiv_closed,
|
||||||
{ apply equiv.to_is_equiv, apply is_contr.sigma_char},
|
{ apply equiv.to_is_equiv, apply is_contr.sigma_char},
|
||||||
apply is_hprop.mk, intros,
|
apply is_prop.mk, intros,
|
||||||
fapply sigma_eq, apply x.2,
|
fapply sigma_eq, apply x.2,
|
||||||
apply is_hprop.elimo},
|
apply is_prop.elimo},
|
||||||
{ intro A,
|
{ intro A,
|
||||||
apply is_trunc_is_equiv_closed,
|
apply is_trunc_is_equiv_closed,
|
||||||
apply equiv.to_is_equiv,
|
apply equiv.to_is_equiv,
|
||||||
|
|
|
@ -37,11 +37,11 @@ section
|
||||||
eq.rec_on (f (refl x)) rfl,
|
eq.rec_on (f (refl x)) rfl,
|
||||||
eq.rec_on p aux
|
eq.rec_on p aux
|
||||||
|
|
||||||
definition is_hset_of_decidable_eq : is_hset A :=
|
definition is_set_of_decidable_eq : is_set A :=
|
||||||
is_hset.mk A (λ x y p q, calc
|
is_set.mk A (λ x y p q, calc
|
||||||
p = (f (refl x))⁻¹ ⬝ (f p) : aux
|
p = (f (refl x))⁻¹ ⬝ (f p) : aux
|
||||||
... = (f (refl x))⁻¹ ⬝ (f q) : f_const
|
... = (f (refl x))⁻¹ ⬝ (f q) : f_const
|
||||||
... = q : aux)
|
... = q : aux)
|
||||||
end
|
end
|
||||||
|
|
||||||
attribute is_hset_of_decidable_eq [instance] [priority 600]
|
attribute is_set_of_decidable_eq [instance] [priority 600]
|
||||||
|
|
|
@ -108,8 +108,8 @@ open nat num is_trunc.trunc_index
|
||||||
namespace is_trunc
|
namespace is_trunc
|
||||||
|
|
||||||
abbreviation is_contr := is_trunc -2
|
abbreviation is_contr := is_trunc -2
|
||||||
abbreviation is_hprop := is_trunc -1
|
abbreviation is_prop := is_trunc -1
|
||||||
abbreviation is_hset := is_trunc 0
|
abbreviation is_set := is_trunc 0
|
||||||
|
|
||||||
variables {A B : Type}
|
variables {A B : Type}
|
||||||
|
|
||||||
|
@ -184,41 +184,41 @@ namespace is_trunc
|
||||||
definition is_trunc_of_is_contr (A : Type) (n : trunc_index) [H : is_contr A] : is_trunc n A :=
|
definition is_trunc_of_is_contr (A : Type) (n : trunc_index) [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_hprop (A : Type) (n : trunc_index) [H : is_hprop A]
|
definition is_trunc_succ_of_is_prop (A : Type) (n : trunc_index) [H : is_prop A]
|
||||||
: is_trunc (n.+1) A :=
|
: is_trunc (n.+1) A :=
|
||||||
is_trunc_of_leq A (show -1 ≤ n.+1, from star)
|
is_trunc_of_leq A (show -1 ≤ n.+1, from star)
|
||||||
|
|
||||||
definition is_trunc_succ_succ_of_is_hset (A : Type) (n : trunc_index) [H : is_hset A]
|
definition is_trunc_succ_succ_of_is_set (A : Type) (n : trunc_index) [H : is_set A]
|
||||||
: is_trunc (n.+2) A :=
|
: is_trunc (n.+2) A :=
|
||||||
@(is_trunc_of_leq A (show 0 ≤ n.+2, from proof star qed)) H
|
@(is_trunc_of_leq A (show 0 ≤ n.+2, from proof star qed)) H
|
||||||
|
|
||||||
/- hprops -/
|
/- hprops -/
|
||||||
|
|
||||||
definition is_hprop.elim [H : is_hprop 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_hprop {A : Type} [H : is_hprop A] (x : A) : is_contr A :=
|
definition is_contr_of_inhabited_hprop {A : Type} [H : is_prop A] (x : A) : is_contr A :=
|
||||||
is_contr.mk x (λy, !is_hprop.elim)
|
is_contr.mk x (λy, !is_prop.elim)
|
||||||
|
|
||||||
theorem is_hprop_of_imp_is_contr {A : Type} (H : A → is_contr A) : is_hprop A :=
|
theorem is_prop_of_imp_is_contr {A : Type} (H : A → is_contr A) : is_prop A :=
|
||||||
@is_trunc_succ_intro A -2
|
@is_trunc_succ_intro A -2
|
||||||
(λx y,
|
(λx y,
|
||||||
have H2 [visible] : is_contr A, from H x,
|
have H2 [visible] : is_contr A, from H x,
|
||||||
!is_contr_eq)
|
!is_contr_eq)
|
||||||
|
|
||||||
theorem is_hprop.mk {A : Type} (H : ∀x y : A, x = y) : is_hprop A :=
|
theorem is_prop.mk {A : Type} (H : ∀x y : A, x = y) : is_prop A :=
|
||||||
is_hprop_of_imp_is_contr (λ x, is_contr.mk x (H x))
|
is_prop_of_imp_is_contr (λ x, is_contr.mk x (H x))
|
||||||
|
|
||||||
theorem is_hprop_elim_self {A : Type} {H : is_hprop A} (x : A) : is_hprop.elim x x = idp :=
|
theorem is_prop_elim_self {A : Type} {H : is_prop A} (x : A) : is_prop.elim x x = idp :=
|
||||||
!is_hprop.elim
|
!is_prop.elim
|
||||||
|
|
||||||
/- hsets -/
|
/- hsets -/
|
||||||
|
|
||||||
theorem is_hset.mk (A : Type) (H : ∀(x y : A) (p q : x = y), p = q) : is_hset A :=
|
theorem is_set.mk (A : Type) (H : ∀(x y : A) (p q : x = y), p = q) : is_set A :=
|
||||||
@is_trunc_succ_intro _ _ (λ x y, is_hprop.mk (H x y))
|
@is_trunc_succ_intro _ _ (λ x y, is_prop.mk (H x y))
|
||||||
|
|
||||||
definition is_hset.elim [H : is_hset A] ⦃x y : A⦄ (p q : x = y) : p = q :=
|
definition is_set.elim [H : is_set A] ⦃x y : A⦄ (p q : x = y) : p = q :=
|
||||||
!is_hprop.elim
|
!is_prop.elim
|
||||||
|
|
||||||
/- instances -/
|
/- instances -/
|
||||||
|
|
||||||
|
@ -241,16 +241,16 @@ namespace is_trunc
|
||||||
definition is_contr_unit : is_contr unit :=
|
definition is_contr_unit : is_contr unit :=
|
||||||
is_contr.mk star (λp, unit.rec_on p idp)
|
is_contr.mk star (λp, unit.rec_on p idp)
|
||||||
|
|
||||||
definition is_hprop_empty : is_hprop empty :=
|
definition is_prop_empty : is_prop empty :=
|
||||||
is_hprop.mk (λx, !empty.elim x)
|
is_prop.mk (λx, !empty.elim x)
|
||||||
|
|
||||||
local attribute is_contr_unit is_hprop_empty [instance]
|
local attribute is_contr_unit is_prop_empty [instance]
|
||||||
|
|
||||||
definition is_trunc_unit [instance] (n : trunc_index) : is_trunc n unit :=
|
definition is_trunc_unit [instance] (n : trunc_index) : is_trunc n unit :=
|
||||||
!is_trunc_of_is_contr
|
!is_trunc_of_is_contr
|
||||||
|
|
||||||
definition is_trunc_empty [instance] (n : trunc_index) : is_trunc (n.+1) empty :=
|
definition is_trunc_empty [instance] (n : trunc_index) : is_trunc (n.+1) empty :=
|
||||||
!is_trunc_succ_of_is_hprop
|
!is_trunc_succ_of_is_prop
|
||||||
|
|
||||||
/- interaction with equivalences -/
|
/- interaction with equivalences -/
|
||||||
|
|
||||||
|
@ -289,16 +289,16 @@ namespace is_trunc
|
||||||
: is_trunc n A :=
|
: 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_hprop [constructor] [HA : is_hprop A] [HB : is_hprop B]
|
definition is_equiv_of_is_prop [constructor] [HA : is_prop A] [HB : is_prop B]
|
||||||
(f : A → B) (g : B → A) : is_equiv f :=
|
(f : A → B) (g : B → A) : is_equiv f :=
|
||||||
is_equiv.mk f g (λb, !is_hprop.elim) (λa, !is_hprop.elim) (λa, !is_hset.elim)
|
is_equiv.mk f g (λb, !is_prop.elim) (λa, !is_prop.elim) (λa, !is_set.elim)
|
||||||
|
|
||||||
definition equiv_of_is_hprop [constructor] [HA : is_hprop A] [HB : is_hprop B]
|
definition equiv_of_is_prop [constructor] [HA : is_prop A] [HB : is_prop B]
|
||||||
(f : A → B) (g : B → A) : A ≃ B :=
|
(f : A → B) (g : B → A) : A ≃ B :=
|
||||||
equiv.mk f (is_equiv_of_is_hprop f g)
|
equiv.mk f (is_equiv_of_is_prop f g)
|
||||||
|
|
||||||
definition equiv_of_iff_of_is_hprop [unfold 5] [HA : is_hprop A] [HB : is_hprop B] (H : A ↔ B) : A ≃ B :=
|
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_hprop (iff.elim_left H) (iff.elim_right H)
|
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 : trunc_index)
|
definition is_trunc_lift [instance] [priority 1450] (A : Type) (n : trunc_index)
|
||||||
|
@ -324,16 +324,16 @@ namespace is_trunc
|
||||||
{a a₂ : A} (p : a = a₂)
|
{a a₂ : A} (p : a = a₂)
|
||||||
(c : C a) (c₂ : C a₂)
|
(c : C a) (c₂ : C a₂)
|
||||||
|
|
||||||
definition is_hprop.elimo [H : is_hprop (C a)] : c =[p] c₂ :=
|
definition is_prop.elimo [H : is_prop (C a)] : c =[p] c₂ :=
|
||||||
pathover_of_eq_tr !is_hprop.elim
|
pathover_of_eq_tr !is_prop.elim
|
||||||
|
|
||||||
definition is_trunc_pathover [instance]
|
definition is_trunc_pathover [instance]
|
||||||
(n : trunc_index) [H : is_trunc (n.+1) (C a)] : is_trunc n (c =[p] c₂) :=
|
(n : trunc_index) [H : is_trunc (n.+1) (C a)] : is_trunc n (c =[p] c₂) :=
|
||||||
is_trunc_equiv_closed_rev n !pathover_equiv_eq_tr
|
is_trunc_equiv_closed_rev n !pathover_equiv_eq_tr
|
||||||
|
|
||||||
variables {p c c₂}
|
variables {p c c₂}
|
||||||
theorem is_hset.elimo (q q' : c =[p] c₂) [H : is_hset (C a)] : q = q' :=
|
theorem is_set.elimo (q q' : c =[p] c₂) [H : is_set (C a)] : q = q' :=
|
||||||
!is_hprop.elim
|
!is_prop.elim
|
||||||
|
|
||||||
-- TODO: port "Truncated morphisms"
|
-- TODO: port "Truncated morphisms"
|
||||||
|
|
||||||
|
|
|
@ -11,6 +11,6 @@ import init.trunc
|
||||||
|
|
||||||
open is_trunc
|
open is_trunc
|
||||||
|
|
||||||
definition eq_rec_eq.{l₁ l₂} {A : Type.{l₁}} {B : A → Type.{l₂}} [h : is_hset A] {a : A} (b : B a) (p : a = a) :
|
definition eq_rec_eq.{l₁ l₂} {A : Type.{l₁}} {B : A → Type.{l₂}} [h : is_set A] {a : A} (b : B a) (p : a = a) :
|
||||||
b = @eq.rec.{l₂ l₁} A a (λ (a' : A) (h : a = a'), B a') b a p :=
|
b = @eq.rec.{l₂ l₁} A a (λ (a' : A) (h : a = a'), B a') b a p :=
|
||||||
eq.rec_on (is_hset.elim (eq.refl a) p) (eq.refl (eq.rec_on (eq.refl a) b))
|
eq.rec_on (is_set.elim (eq.refl a) p) (eq.refl (eq.rec_on (eq.refl a) b))
|
||||||
|
|
|
@ -9,10 +9,10 @@ import types.pi
|
||||||
|
|
||||||
open eq is_trunc decidable
|
open eq is_trunc decidable
|
||||||
|
|
||||||
theorem dif_pos {c : Type} [H : decidable c] [P : is_hprop c] (Hc : c)
|
theorem dif_pos {c : Type} [H : decidable c] [P : is_prop c] (Hc : c)
|
||||||
{A : Type} {t : c → A} {e : ¬ c → A} : dite c t e = t Hc :=
|
{A : Type} {t : c → A} {e : ¬ c → A} : dite c t e = t Hc :=
|
||||||
by induction H with Hc Hnc; apply ap t; apply is_hprop.elim; apply absurd Hc Hnc
|
by induction H with Hc Hnc; apply ap t; apply is_prop.elim; apply absurd Hc Hnc
|
||||||
|
|
||||||
theorem dif_neg {c : Type} [H : decidable c] (Hnc : ¬c)
|
theorem dif_neg {c : Type} [H : decidable c] (Hnc : ¬c)
|
||||||
{A : Type} {t : c → A} {e : ¬ c → A} : dite c t e = e Hnc :=
|
{A : Type} {t : c → A} {e : ¬ c → A} : dite c t e = e Hnc :=
|
||||||
by induction H with Hc Hnc; apply absurd Hc Hnc; apply ap e; apply is_hprop.elim
|
by induction H with Hc Hnc; apply absurd Hc Hnc; apply ap e; apply is_prop.elim
|
||||||
|
|
|
@ -155,9 +155,9 @@ namespace bool
|
||||||
assert H2 : bnot = id, from !cast_ua_fn⁻¹ ⬝ ap cast H,
|
assert H2 : bnot = id, from !cast_ua_fn⁻¹ ⬝ ap cast H,
|
||||||
absurd (ap10 H2 tt) ff_ne_tt
|
absurd (ap10 H2 tt) ff_ne_tt
|
||||||
|
|
||||||
theorem is_hset_bool : is_hset bool := _
|
theorem is_set_bool : is_set bool := _
|
||||||
theorem not_is_hprop_bool_eq_bool : ¬ is_hprop (bool = bool) :=
|
theorem not_is_prop_bool_eq_bool : ¬ is_prop (bool = bool) :=
|
||||||
λ H, eq_bnot_ne_idp !is_hprop.elim
|
λ H, eq_bnot_ne_idp !is_prop.elim
|
||||||
|
|
||||||
definition bool_equiv_option_unit [constructor] : bool ≃ option unit :=
|
definition bool_equiv_option_unit [constructor] : bool ≃ option unit :=
|
||||||
begin
|
begin
|
||||||
|
|
|
@ -460,7 +460,7 @@ namespace eq
|
||||||
(p ⬝ q) ▸ (center (Σa, code a)).2
|
(p ⬝ q) ▸ (center (Σa, code a)).2
|
||||||
|
|
||||||
definition decode' {a : A} (c : code a) : a₀ = a :=
|
definition decode' {a : A} (c : code a) : a₀ = a :=
|
||||||
(is_hprop.elim ⟨a₀, encode idp⟩ ⟨a, c⟩)..1
|
(is_prop.elim ⟨a₀, encode idp⟩ ⟨a, c⟩)..1
|
||||||
|
|
||||||
definition decode {a : A} (c : code a) : a₀ = a :=
|
definition decode {a : A} (c : code a) : a₀ = a :=
|
||||||
(decode' (encode idp))⁻¹ ⬝ decode' c
|
(decode' (encode idp))⁻¹ ⬝ decode' c
|
||||||
|
@ -472,8 +472,8 @@ namespace eq
|
||||||
{ exact decode},
|
{ exact decode},
|
||||||
{ intro c,
|
{ intro c,
|
||||||
unfold [encode, decode, decode'],
|
unfold [encode, decode, decode'],
|
||||||
induction p, esimp, rewrite [is_hprop_elim_self,▸*,+idp_con], apply tr_eq_of_pathover,
|
induction p, esimp, rewrite [is_prop_elim_self,▸*,+idp_con], apply tr_eq_of_pathover,
|
||||||
eapply @sigma.rec_on _ _ (λx, x.2 =[(is_hprop.elim ⟨x.1, x.2⟩ ⟨a, c⟩)..1] c)
|
eapply @sigma.rec_on _ _ (λx, x.2 =[(is_prop.elim ⟨x.1, x.2⟩ ⟨a, c⟩)..1] c)
|
||||||
(center (sigma code)), -- BUG(?): induction fails
|
(center (sigma code)), -- BUG(?): induction fails
|
||||||
intro a c, apply eq_pr2},
|
intro a c, apply eq_pr2},
|
||||||
{ intro q, induction q, esimp, apply con.left_inv, },
|
{ intro q, induction q, esimp, apply con.left_inv, },
|
||||||
|
|
|
@ -75,19 +75,19 @@ namespace is_equiv
|
||||||
local attribute is_contr_right_inverse [instance] [priority 1600]
|
local attribute is_contr_right_inverse [instance] [priority 1600]
|
||||||
local attribute is_contr_right_coherence [instance] [priority 1600]
|
local attribute is_contr_right_coherence [instance] [priority 1600]
|
||||||
|
|
||||||
theorem is_hprop_is_equiv [instance] : is_hprop (is_equiv f) :=
|
theorem is_prop_is_equiv [instance] : is_prop (is_equiv f) :=
|
||||||
is_hprop_of_imp_is_contr
|
is_prop_of_imp_is_contr
|
||||||
(λ(H : is_equiv f), is_trunc_equiv_closed -2 (equiv.symm !is_equiv.sigma_char'))
|
(λ(H : is_equiv f), is_trunc_equiv_closed -2 (equiv.symm !is_equiv.sigma_char'))
|
||||||
|
|
||||||
definition inv_eq_inv {A B : Type} {f f' : A → B} {Hf : is_equiv f} {Hf' : is_equiv f'}
|
definition inv_eq_inv {A B : Type} {f f' : A → B} {Hf : is_equiv f} {Hf' : is_equiv f'}
|
||||||
(p : f = f') : f⁻¹ = f'⁻¹ :=
|
(p : f = f') : f⁻¹ = f'⁻¹ :=
|
||||||
apd011 inv p !is_hprop.elim
|
apd011 inv p !is_prop.elim
|
||||||
|
|
||||||
/- contractible fibers -/
|
/- contractible fibers -/
|
||||||
definition is_contr_fun_of_is_equiv [H : is_equiv f] : is_contr_fun f :=
|
definition is_contr_fun_of_is_equiv [H : is_equiv f] : is_contr_fun f :=
|
||||||
is_contr_fiber_of_is_equiv f
|
is_contr_fiber_of_is_equiv f
|
||||||
|
|
||||||
definition is_hprop_is_contr_fun (f : A → B) : is_hprop (is_contr_fun f) := _
|
definition is_prop_is_contr_fun (f : A → B) : is_prop (is_contr_fun f) := _
|
||||||
|
|
||||||
definition is_equiv_of_is_contr_fun [H : is_contr_fun f] : is_equiv f :=
|
definition is_equiv_of_is_contr_fun [H : is_contr_fun f] : is_equiv f :=
|
||||||
adjointify _ (λb, point (center (fiber f b)))
|
adjointify _ (λb, point (center (fiber f b)))
|
||||||
|
@ -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_hprop _ (λH, !is_equiv_of_is_contr_fun)
|
equiv_of_is_prop _ (λH, !is_equiv_of_is_contr_fun)
|
||||||
|
|
||||||
end is_equiv
|
end is_equiv
|
||||||
|
|
||||||
|
@ -129,7 +129,7 @@ namespace equiv
|
||||||
|
|
||||||
definition equiv_mk_eq {f f' : A → B} [H : is_equiv f] [H' : is_equiv f'] (p : f = f')
|
definition equiv_mk_eq {f f' : A → B} [H : is_equiv f] [H' : is_equiv f'] (p : f = f')
|
||||||
: equiv.mk f H = equiv.mk f' H' :=
|
: equiv.mk f H = equiv.mk f' H' :=
|
||||||
apd011 equiv.mk p !is_hprop.elim
|
apd011 equiv.mk p !is_prop.elim
|
||||||
|
|
||||||
definition equiv_eq {f f' : A ≃ B} (p : to_fun f = to_fun f') : f = f' :=
|
definition equiv_eq {f f' : A ≃ B} (p : to_fun f = to_fun f') : f = f' :=
|
||||||
by (cases f; cases f'; apply (equiv_mk_eq p))
|
by (cases f; cases f'; apply (equiv_mk_eq p))
|
||||||
|
@ -164,11 +164,11 @@ namespace equiv
|
||||||
: is_equiv (ap to_fun : f = f' → to_fun f = to_fun f') :=
|
: is_equiv (ap to_fun : f = f' → to_fun f = to_fun f') :=
|
||||||
begin
|
begin
|
||||||
fapply adjointify,
|
fapply adjointify,
|
||||||
{intro p, cases f with f H, cases f' with f' H', cases p, apply ap (mk f'), apply is_hprop.elim},
|
{intro p, cases f with f H, cases f' with f' H', cases p, apply ap (mk f'), apply is_prop.elim},
|
||||||
{intro p, cases f with f H, cases f' with f' H', cases p,
|
{intro p, cases f with f H, cases f' with f' H', cases p,
|
||||||
apply @concat _ _ (ap to_fun (ap (equiv.mk f') (is_hprop.elim H H'))), {apply idp},
|
apply @concat _ _ (ap to_fun (ap (equiv.mk f') (is_prop.elim H H'))), {apply idp},
|
||||||
generalize is_hprop.elim H H', intro q, cases q, apply idp},
|
generalize is_prop.elim H H', intro q, cases q, apply idp},
|
||||||
{intro p, cases p, cases f with f H, apply ap (ap (equiv.mk f)), apply is_hset.elim}
|
{intro p, cases p, cases f with f H, apply ap (ap (equiv.mk f)), apply is_set.elim}
|
||||||
end
|
end
|
||||||
|
|
||||||
definition equiv_pathover {A : Type} {a a' : A} (p : a = a')
|
definition equiv_pathover {A : Type} {a a' : A} (p : a = a')
|
||||||
|
@ -179,21 +179,21 @@ namespace equiv
|
||||||
{ intro a, apply equiv.sigma_char},
|
{ intro a, apply equiv.sigma_char},
|
||||||
{ fapply sigma_pathover,
|
{ fapply sigma_pathover,
|
||||||
esimp, apply arrow_pathover, exact r,
|
esimp, apply arrow_pathover, exact r,
|
||||||
apply is_hprop.elimo}
|
apply is_prop.elimo}
|
||||||
end
|
end
|
||||||
|
|
||||||
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_hprop, apply is_hprop.mk,
|
apply @is_contr_of_inhabited_hprop, apply is_prop.mk,
|
||||||
intro x y, cases x with fx Hx, cases y with fy Hy, generalize Hy,
|
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_hprop.elim ▸ rfl),
|
apply (eq_of_homotopy (λ a, !eq_of_is_contr)) ▸ (λ Hy, !is_prop.elim ▸ rfl),
|
||||||
apply equiv_of_is_contr_of_is_contr
|
apply equiv_of_is_contr_of_is_contr
|
||||||
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_hprop))
|
(@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) :=
|
||||||
|
|
|
@ -105,7 +105,7 @@ namespace fiber
|
||||||
{ intro a, apply fiber.mk a, reflexivity },
|
{ intro a, apply fiber.mk a, reflexivity },
|
||||||
{ intro a, reflexivity },
|
{ intro a, reflexivity },
|
||||||
{ intro f, cases f with a H, change fiber.mk a (refl star) = fiber.mk a H,
|
{ intro f, cases f with a H, change fiber.mk a (refl star) = fiber.mk a H,
|
||||||
rewrite [is_hset.elim H (refl star)] }
|
rewrite [is_set.elim H (refl star)] }
|
||||||
end
|
end
|
||||||
|
|
||||||
definition fiber_const_equiv (A : Type) (a₀ : A) (a : A)
|
definition fiber_const_equiv (A : Type) (a₀ : A) (a : A)
|
||||||
|
|
|
@ -28,7 +28,7 @@ begin
|
||||||
intro i, cases i with i ilt, reflexivity
|
intro i, cases i with i ilt, reflexivity
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_hset_fin [instance] : is_hset (fin n) :=
|
definition is_set_fin [instance] : is_set (fin n) :=
|
||||||
begin
|
begin
|
||||||
apply is_trunc_equiv_closed, apply equiv.symm, apply sigma_char,
|
apply is_trunc_equiv_closed, apply equiv.symm, apply sigma_char,
|
||||||
end
|
end
|
||||||
|
@ -36,11 +36,11 @@ end
|
||||||
definition eq_of_veq : Π {i j : fin n}, (val i) = j → i = j :=
|
definition eq_of_veq : Π {i j : fin n}, (val i) = j → i = j :=
|
||||||
begin
|
begin
|
||||||
intro i j, cases i with i ilt, cases j with j jlt, esimp,
|
intro i j, cases i with i ilt, cases j with j jlt, esimp,
|
||||||
intro p, induction p, apply ap (mk i), apply !is_hprop.elim
|
intro p, induction p, apply ap (mk i), apply !is_prop.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
definition eq_of_veq_refl (i : fin n) : eq_of_veq (refl (val i)) = idp :=
|
definition eq_of_veq_refl (i : fin n) : eq_of_veq (refl (val i)) = idp :=
|
||||||
!is_hprop.elim
|
!is_prop.elim
|
||||||
|
|
||||||
definition veq_of_eq : Π {i j : fin n}, i = j → (val i) = j :=
|
definition veq_of_eq : Π {i j : fin n}, i = j → (val i) = j :=
|
||||||
by intro i j P; apply ap val; exact P
|
by intro i j P; apply ap val; exact P
|
||||||
|
@ -379,9 +379,9 @@ begin
|
||||||
apply decidable.rec,
|
apply decidable.rec,
|
||||||
{ intro ilt', esimp[val_inj], apply concat,
|
{ intro ilt', esimp[val_inj], apply concat,
|
||||||
apply ap (λ x, eq.rec_on x _), esimp[eq_of_veq, rfl], reflexivity,
|
apply ap (λ x, eq.rec_on x _), esimp[eq_of_veq, rfl], reflexivity,
|
||||||
assert H : ilt = ilt', apply is_hprop.elim, cases H,
|
assert H : ilt = ilt', apply is_prop.elim, cases H,
|
||||||
assert H' : is_hprop.elim (lt_add_of_lt_right ilt 1) (lt_add_of_lt_right ilt 1) = idp,
|
assert H' : is_prop.elim (lt_add_of_lt_right ilt 1) (lt_add_of_lt_right ilt 1) = idp,
|
||||||
apply is_hprop.elim,
|
apply is_prop.elim,
|
||||||
krewrite H' },
|
krewrite H' },
|
||||||
{ intro a, contradiction },
|
{ intro a, contradiction },
|
||||||
end
|
end
|
||||||
|
@ -395,7 +395,7 @@ begin
|
||||||
{ intro a, apply absurd a !nat.lt_irrefl },
|
{ intro a, apply absurd a !nat.lt_irrefl },
|
||||||
{ intro a, esimp[val_inj], apply concat,
|
{ intro a, esimp[val_inj], apply concat,
|
||||||
assert H : (le.antisymm (le_of_lt_succ (lt_succ_self n)) (le_of_not_gt a))⁻¹ = idp,
|
assert H : (le.antisymm (le_of_lt_succ (lt_succ_self n)) (le_of_not_gt a))⁻¹ = idp,
|
||||||
apply is_hprop.elim,
|
apply is_prop.elim,
|
||||||
apply ap _ H, krewrite eq_of_veq_refl },
|
apply ap _ H, krewrite eq_of_veq_refl },
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -573,7 +573,7 @@ protected definition integral_domain [reducible] [trans_instance] : integral_dom
|
||||||
mul_comm := int.mul_comm,
|
mul_comm := int.mul_comm,
|
||||||
zero_ne_one := int.zero_ne_one,
|
zero_ne_one := int.zero_ne_one,
|
||||||
eq_zero_sum_eq_zero_of_mul_eq_zero := @int.eq_zero_sum_eq_zero_of_mul_eq_zero,
|
eq_zero_sum_eq_zero_of_mul_eq_zero := @int.eq_zero_sum_eq_zero_of_mul_eq_zero,
|
||||||
is_hset_carrier := is_hset_of_decidable_eq⦄
|
is_set_carrier := is_set_of_decidable_eq⦄
|
||||||
|
|
||||||
definition int_has_sub [reducible] [instance] [priority int.prio] : has_sub int :=
|
definition int_has_sub [reducible] [instance] [priority int.prio] : has_sub int :=
|
||||||
has_sub.mk has_sub.sub
|
has_sub.mk has_sub.sub
|
||||||
|
|
|
@ -137,7 +137,7 @@ rfl
|
||||||
|
|
||||||
theorem last_congr {l₁ l₂ : list T} (h₁ : l₁ ≠ []) (h₂ : l₂ ≠ []) (h₃ : l₁ = l₂)
|
theorem last_congr {l₁ l₂ : list T} (h₁ : l₁ ≠ []) (h₂ : l₂ ≠ []) (h₃ : l₁ = l₂)
|
||||||
: last l₁ h₁ = last l₂ h₂ :=
|
: last l₁ h₁ = last l₂ h₂ :=
|
||||||
apd011 last h₃ !is_hprop.elim
|
apd011 last h₃ !is_prop.elim
|
||||||
|
|
||||||
theorem last_concat {x : T} : ∀ {l : list T} (h : concat x l ≠ []), last (concat x l) h = x
|
theorem last_concat {x : T} : ∀ {l : list T} (h : concat x l ≠ []), last (concat x l) h = x
|
||||||
| [] h := rfl
|
| [] h := rfl
|
||||||
|
|
|
@ -305,7 +305,7 @@ protected definition comm_semiring [reducible] [trans_instance] : comm_semiring
|
||||||
zero_mul := nat.zero_mul,
|
zero_mul := nat.zero_mul,
|
||||||
mul_zero := nat.mul_zero,
|
mul_zero := nat.mul_zero,
|
||||||
mul_comm := nat.mul_comm,
|
mul_comm := nat.mul_comm,
|
||||||
is_hset_carrier:= _⦄
|
is_set_carrier:= _⦄
|
||||||
end nat
|
end nat
|
||||||
|
|
||||||
section
|
section
|
||||||
|
|
|
@ -11,32 +11,32 @@ import .order
|
||||||
open is_trunc unit empty eq equiv algebra
|
open is_trunc unit empty eq equiv algebra
|
||||||
|
|
||||||
namespace nat
|
namespace nat
|
||||||
definition is_hprop_le [instance] (n m : ℕ) : is_hprop (n ≤ m) :=
|
definition is_prop_le [instance] (n m : ℕ) : is_prop (n ≤ m) :=
|
||||||
begin
|
begin
|
||||||
assert lem : Π{n m : ℕ} (p : n ≤ m) (q : n = m), p = q ▸ le.refl n,
|
assert lem : Π{n m : ℕ} (p : n ≤ m) (q : n = m), p = q ▸ le.refl n,
|
||||||
{ intros, cases p,
|
{ intros, cases p,
|
||||||
{ assert H' : q = idp, apply is_hset.elim,
|
{ assert H' : q = idp, apply is_set.elim,
|
||||||
cases H', reflexivity},
|
cases H', reflexivity},
|
||||||
{ cases q, exfalso, apply not_succ_le_self a}},
|
{ cases q, exfalso, apply not_succ_le_self a}},
|
||||||
apply is_hprop.mk, intro H1 H2, induction H2,
|
apply is_prop.mk, intro H1 H2, induction H2,
|
||||||
{ apply lem},
|
{ apply lem},
|
||||||
{ cases H1,
|
{ cases H1,
|
||||||
{ exfalso, apply not_succ_le_self a},
|
{ exfalso, apply not_succ_le_self a},
|
||||||
{ exact ap le.step !v_0}},
|
{ exact ap le.step !v_0}},
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_hprop_lt [instance] (n m : ℕ) : is_hprop (n < m) := !is_hprop_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_hprop 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_hprop 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 :=
|
||||||
begin
|
begin
|
||||||
unfold lt.by_cases, induction (lt.trichotomy a b) with H' H',
|
unfold lt.by_cases, induction (lt.trichotomy a b) with H' H',
|
||||||
{ esimp, exact ap H1 !is_hprop.elim},
|
{ esimp, exact ap H1 !is_prop.elim},
|
||||||
{ exfalso, cases H' with H' H', apply lt.irrefl, exact H' ▸ H, exact lt.asymm H H'}
|
{ exfalso, cases H' with H' H', apply lt.irrefl, exact H' ▸ H, exact lt.asymm H H'}
|
||||||
end
|
end
|
||||||
|
|
||||||
|
@ -45,7 +45,7 @@ namespace nat
|
||||||
begin
|
begin
|
||||||
unfold lt.by_cases, induction (lt.trichotomy a b) with H' H',
|
unfold lt.by_cases, induction (lt.trichotomy a b) with H' H',
|
||||||
{ exfalso, apply lt.irrefl, exact H ▸ H'},
|
{ exfalso, apply lt.irrefl, exact H ▸ H'},
|
||||||
{ cases H' with H' H', esimp, exact ap H2 !is_hprop.elim, exfalso, apply lt.irrefl, exact H ▸ H'}
|
{ cases H' with H' H', esimp, exact ap H2 !is_prop.elim, exfalso, apply lt.irrefl, exact H ▸ H'}
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem lt_by_cases_ge {a b : ℕ} {P : Type} (H1 : a < b → P) (H2 : a = b → P)
|
theorem lt_by_cases_ge {a b : ℕ} {P : Type} (H1 : a < b → P) (H2 : a = b → P)
|
||||||
|
@ -53,7 +53,7 @@ namespace nat
|
||||||
begin
|
begin
|
||||||
unfold lt.by_cases, induction (lt.trichotomy a b) with H' H',
|
unfold lt.by_cases, induction (lt.trichotomy a b) with H' H',
|
||||||
{ exfalso, exact lt.asymm H H'},
|
{ exfalso, exact lt.asymm H H'},
|
||||||
{ cases H' with H' H', exfalso, apply lt.irrefl, exact H' ▸ H, esimp, exact ap H3 !is_hprop.elim}
|
{ cases H' with H' H', exfalso, apply lt.irrefl, exact H' ▸ H, esimp, exact ap H3 !is_prop.elim}
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem lt_ge_by_cases_lt {n m : ℕ} {P : Type} (H1 : n < m → P) (H2 : n ≥ m → P)
|
theorem lt_ge_by_cases_lt {n m : ℕ} {P : Type} (H1 : n < m → P) (H2 : n ≥ m → P)
|
||||||
|
@ -65,7 +65,7 @@ namespace nat
|
||||||
begin
|
begin
|
||||||
unfold [lt_ge_by_cases,lt.by_cases], induction (lt.trichotomy n m) with H' H',
|
unfold [lt_ge_by_cases,lt.by_cases], induction (lt.trichotomy n m) with H' H',
|
||||||
{ exfalso, apply lt.irrefl, exact lt_of_le_of_lt H H'},
|
{ exfalso, apply lt.irrefl, exact lt_of_le_of_lt H H'},
|
||||||
{ cases H' with H' H'; all_goals (esimp; apply ap H2 !is_hprop.elim)}
|
{ cases H' with H' H'; all_goals (esimp; apply ap H2 !is_prop.elim)}
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem lt_ge_by_cases_le {n m : ℕ} {P : Type} {H1 : n ≤ m → P} {H2 : n ≥ m → P}
|
theorem lt_ge_by_cases_le {n m : ℕ} {P : Type} {H1 : n ≤ m → P} {H2 : n ≥ m → P}
|
||||||
|
@ -73,10 +73,10 @@ namespace nat
|
||||||
: lt_ge_by_cases (λH', H1 (le_of_lt H')) H2 = H1 H :=
|
: lt_ge_by_cases (λH', H1 (le_of_lt H')) H2 = H1 H :=
|
||||||
begin
|
begin
|
||||||
unfold [lt_ge_by_cases,lt.by_cases], induction (lt.trichotomy n m) with H' H',
|
unfold [lt_ge_by_cases,lt.by_cases], induction (lt.trichotomy n m) with H' H',
|
||||||
{ esimp, apply ap H1 !is_hprop.elim},
|
{ esimp, apply ap H1 !is_prop.elim},
|
||||||
{ cases H' with H' H',
|
{ cases H' with H' H',
|
||||||
{ esimp, induction H', esimp, symmetry,
|
{ esimp, induction H', esimp, symmetry,
|
||||||
exact ap H1 !is_hprop.elim ⬝ Heq idp ⬝ ap H2 !is_hprop.elim},
|
exact ap H1 !is_prop.elim ⬝ Heq idp ⬝ ap H2 !is_prop.elim},
|
||||||
{ exfalso, apply lt.irrefl, apply lt_of_le_of_lt H H'}}
|
{ exfalso, apply lt.irrefl, apply lt_of_le_of_lt H H'}}
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -238,7 +238,7 @@ namespace pi
|
||||||
{ intro f, exact star},
|
{ intro f, exact star},
|
||||||
{ intro u a, exact !center},
|
{ intro u a, exact !center},
|
||||||
{ intro u, induction u, reflexivity},
|
{ intro u, induction u, reflexivity},
|
||||||
{ intro f, apply eq_of_homotopy, intro a, apply is_hprop.elim}
|
{ intro f, apply eq_of_homotopy, intro a, apply is_prop.elim}
|
||||||
end
|
end
|
||||||
|
|
||||||
/- Interaction with other type constructors -/
|
/- Interaction with other type constructors -/
|
||||||
|
@ -298,15 +298,15 @@ namespace pi
|
||||||
theorem is_trunc_not [instance] (n : trunc_index) (A : Type) : is_trunc (n.+1) ¬A :=
|
theorem is_trunc_not [instance] (n : trunc_index) (A : Type) : is_trunc (n.+1) ¬A :=
|
||||||
by unfold not;exact _
|
by unfold not;exact _
|
||||||
|
|
||||||
theorem is_hprop_pi_eq [instance] [priority 490] (a : A) : is_hprop (Π(a' : A), a = a') :=
|
theorem is_prop_pi_eq [instance] [priority 490] (a : A) : is_prop (Π(a' : A), a = a') :=
|
||||||
is_hprop_of_imp_is_contr
|
is_prop_of_imp_is_contr
|
||||||
( assume (f : Πa', a = a'),
|
( assume (f : Πa', a = a'),
|
||||||
assert is_contr A, from is_contr.mk a f,
|
assert is_contr A, from is_contr.mk a f,
|
||||||
by exact _) /- force type clas resolution -/
|
by exact _) /- force type clas resolution -/
|
||||||
|
|
||||||
theorem is_hprop_neg (A : Type) : is_hprop (¬A) := _
|
theorem is_prop_neg (A : Type) : is_prop (¬A) := _
|
||||||
local attribute ne [reducible]
|
local attribute ne [reducible]
|
||||||
theorem is_hprop_ne [instance] {A : Type} (a b : A) : is_hprop (a ≠ b) := _
|
theorem is_prop_ne [instance] {A : Type} (a b : A) : is_prop (a ≠ b) := _
|
||||||
|
|
||||||
/- Symmetry of Π -/
|
/- Symmetry of Π -/
|
||||||
definition is_equiv_flip [instance] {P : A → A' → Type}
|
definition is_equiv_flip [instance] {P : A → A' → Type}
|
||||||
|
|
|
@ -87,7 +87,7 @@ namespace pointed
|
||||||
definition pequiv_eq {p q : A ≃* B} (H : p = q :> (A →* B)) : p = q :=
|
definition pequiv_eq {p q : A ≃* B} (H : p = q :> (A →* B)) : p = q :=
|
||||||
begin
|
begin
|
||||||
cases p with f Hf, cases q with g Hg, esimp at *,
|
cases p with f Hf, cases q with g Hg, esimp at *,
|
||||||
exact apd011 pequiv_of_pmap H !is_hprop.elim
|
exact apd011 pequiv_of_pmap H !is_prop.elim
|
||||||
end
|
end
|
||||||
|
|
||||||
definition loop_space_pequiv_rfl
|
definition loop_space_pequiv_rfl
|
||||||
|
|
|
@ -58,7 +58,7 @@ namespace pullback
|
||||||
{ intro v, induction v with x y p, exact (x, y)},
|
{ intro v, induction v with x y p, exact (x, y)},
|
||||||
{ intro v, induction v with x y, exact pullback.mk x y idp},
|
{ intro v, induction v with x y, exact pullback.mk x y idp},
|
||||||
{ intro v, induction v, reflexivity},
|
{ intro v, induction v, reflexivity},
|
||||||
{ intro v, induction v, esimp, apply ap _ !is_hprop.elim},
|
{ intro v, induction v, esimp, apply ap _ !is_prop.elim},
|
||||||
end
|
end
|
||||||
|
|
||||||
definition pullback_along {f : A₂₀ → A₂₂} (g : A₀₂ → A₂₂) : pullback f g → A₂₀ :=
|
definition pullback_along {f : A₂₀ → A₂₂} (g : A₀₂ → A₂₂) : pullback f g → A₂₀ :=
|
||||||
|
|
|
@ -423,28 +423,28 @@ namespace sigma
|
||||||
|
|
||||||
/- Subtypes (sigma types whose second components are hprops) -/
|
/- Subtypes (sigma types whose second components are hprops) -/
|
||||||
|
|
||||||
definition subtype [reducible] {A : Type} (P : A → Type) [H : Πa, is_hprop (P a)] :=
|
definition subtype [reducible] {A : Type} (P : A → Type) [H : Πa, is_prop (P a)] :=
|
||||||
Σ(a : A), P a
|
Σ(a : A), P a
|
||||||
notation [parsing_only] `{` binder `|` r:(scoped:1 P, subtype P) `}` := r
|
notation [parsing_only] `{` binder `|` r:(scoped:1 P, subtype P) `}` := r
|
||||||
|
|
||||||
/- To prove equality in a subtype, we only need equality of the first component. -/
|
/- To prove equality in a subtype, we only need equality of the first component. -/
|
||||||
definition subtype_eq [H : Πa, is_hprop (B a)] {u v : {a | B a}} : u.1 = v.1 → u = v :=
|
definition subtype_eq [H : Πa, is_prop (B a)] {u v : {a | B a}} : u.1 = v.1 → u = v :=
|
||||||
sigma_eq_unc ∘ inv pr1
|
sigma_eq_unc ∘ inv pr1
|
||||||
|
|
||||||
definition is_equiv_subtype_eq [H : Πa, is_hprop (B a)] (u v : {a | B a})
|
definition is_equiv_subtype_eq [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 [H : Πa, is_hprop (B a)] (u v : {a | B a}) : (u.1 = v.1) ≃ (u = v) :=
|
definition equiv_subtype [H : Πa, is_prop (B a)] (u v : {a | B a}) : (u.1 = v.1) ≃ (u = v) :=
|
||||||
equiv.mk !subtype_eq _
|
equiv.mk !subtype_eq _
|
||||||
|
|
||||||
definition subtype_eq_inv {A : Type} {B : A → Type} [H : Πa, is_hprop (B a)] (u v : Σa, B a)
|
definition subtype_eq_inv {A : Type} {B : A → Type} [H : Πa, is_prop (B a)] (u v : Σa, B a)
|
||||||
: u = v → u.1 = v.1 :=
|
: u = v → u.1 = v.1 :=
|
||||||
subtype_eq⁻¹ᶠ
|
subtype_eq⁻¹ᶠ
|
||||||
|
|
||||||
local attribute subtype_eq_inv [reducible]
|
local attribute subtype_eq_inv [reducible]
|
||||||
definition is_equiv_subtype_eq_inv {A : Type} {B : A → Type} [H : Πa, is_hprop (B a)]
|
definition is_equiv_subtype_eq_inv {A : Type} {B : A → Type} [H : Πa, is_prop (B a)]
|
||||||
(u v : Σa, B a) : is_equiv (subtype_eq_inv u v) :=
|
(u v : Σa, B a) : is_equiv (subtype_eq_inv u v) :=
|
||||||
_
|
_
|
||||||
|
|
||||||
|
@ -463,7 +463,7 @@ namespace sigma
|
||||||
|
|
||||||
theorem is_trunc_subtype (B : A → hprop) (n : trunc_index)
|
theorem is_trunc_subtype (B : A → hprop) (n : trunc_index)
|
||||||
[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_hprop)
|
@(is_trunc_sigma B (n.+1)) _ (λa, !is_trunc_succ_of_is_prop)
|
||||||
|
|
||||||
end sigma
|
end sigma
|
||||||
|
|
||||||
|
|
|
@ -215,7 +215,7 @@ namespace sum
|
||||||
cases u' with u' Hu', exfalso, apply empty_of_inl_eq_inr,
|
cases u' with u' Hu', exfalso, apply empty_of_inl_eq_inr,
|
||||||
calc inl ⋆ = H⁻¹ (H (inl ⋆)) : (to_left_inv H (inl ⋆))⁻¹
|
calc inl ⋆ = H⁻¹ (H (inl ⋆)) : (to_left_inv H (inl ⋆))⁻¹
|
||||||
... = H⁻¹ (inl u') : {Hu'}
|
... = H⁻¹ (inl u') : {Hu'}
|
||||||
... = H⁻¹ (inl u) : is_hprop.elim
|
... = H⁻¹ (inl u) : is_prop.elim
|
||||||
... = H⁻¹ (H (inr a)) : {Hu⁻¹}
|
... = H⁻¹ (H (inr a)) : {Hu⁻¹}
|
||||||
... = inr a : to_left_inv H (inr a)
|
... = inr a : to_left_inv H (inr a)
|
||||||
end
|
end
|
||||||
|
@ -229,7 +229,7 @@ namespace sum
|
||||||
{ intro x, exfalso, cases x with u' Hu', apply empty_of_inl_eq_inr,
|
{ intro x, exfalso, cases x with u' Hu', apply empty_of_inl_eq_inr,
|
||||||
calc inl ⋆ = H⁻¹ (H (inl ⋆)) : (to_left_inv H (inl ⋆))⁻¹
|
calc inl ⋆ = H⁻¹ (H (inl ⋆)) : (to_left_inv H (inl ⋆))⁻¹
|
||||||
... = H⁻¹ (inl u') : ap H⁻¹ Hu'
|
... = H⁻¹ (inl u') : ap H⁻¹ Hu'
|
||||||
... = H⁻¹ (inl u) : {!is_hprop.elim}
|
... = H⁻¹ (inl u) : {!is_prop.elim}
|
||||||
... = H⁻¹ (H (inr _)) : {Hu⁻¹}
|
... = H⁻¹ (H (inr _)) : {Hu⁻¹}
|
||||||
... = inr _ : to_left_inv H },
|
... = inr _ : to_left_inv H },
|
||||||
{ intro x, cases x with b' Hb', esimp, cases sum.mem_cases (H⁻¹ (inr b)) with x x,
|
{ intro x, cases x with b' Hb', esimp, cases sum.mem_cases (H⁻¹ (inr b)) with x x,
|
||||||
|
@ -250,7 +250,7 @@ namespace sum
|
||||||
{ cases x with u' Hu', exfalso, apply empty_of_inl_eq_inr,
|
{ cases x with u' Hu', exfalso, apply empty_of_inl_eq_inr,
|
||||||
calc inl ⋆ = H (H⁻¹ (inl ⋆)) : (to_right_inv H (inl ⋆))⁻¹
|
calc inl ⋆ = H (H⁻¹ (inl ⋆)) : (to_right_inv H (inl ⋆))⁻¹
|
||||||
... = H (inl u') : {ap H Hu'}
|
... = H (inl u') : {ap H Hu'}
|
||||||
... = H (inl u) : {!is_hprop.elim}
|
... = H (inl u) : {!is_prop.elim}
|
||||||
... = H (H⁻¹ (inr b)) : {ap H Hu⁻¹}
|
... = H (H⁻¹ (inr b)) : {ap H Hu⁻¹}
|
||||||
... = inr b : to_right_inv H (inr b) },
|
... = inr b : to_right_inv H (inr b) },
|
||||||
{ cases x with a Ha, exfalso, apply empty_of_inl_eq_inr,
|
{ cases x with a Ha, exfalso, apply empty_of_inl_eq_inr,
|
||||||
|
@ -318,9 +318,9 @@ namespace sum
|
||||||
induction n with n IH,
|
induction n with n IH,
|
||||||
{ exfalso, exact H !center !center},
|
{ exfalso, exact H !center !center},
|
||||||
{ clear IH, induction n with n IH,
|
{ clear IH, induction n with n IH,
|
||||||
{ apply is_hprop.mk, intros x y,
|
{ apply is_prop.mk, intros x y,
|
||||||
induction x, all_goals induction y, all_goals esimp,
|
induction x, all_goals induction y, all_goals esimp,
|
||||||
all_goals try (exfalso;apply H;assumption;assumption), all_goals apply ap _ !is_hprop.elim},
|
all_goals try (exfalso;apply H;assumption;assumption), all_goals apply ap _ !is_prop.elim},
|
||||||
{ apply is_trunc_sum}}
|
{ apply is_trunc_sum}}
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -87,13 +87,13 @@ namespace is_trunc
|
||||||
|
|
||||||
|
|
||||||
/- theorems about decidable equality and axiom K -/
|
/- theorems about decidable equality and axiom K -/
|
||||||
theorem is_hset_of_axiom_K {A : Type} (K : Π{a : A} (p : a = a), p = idp) : is_hset A :=
|
theorem is_set_of_axiom_K {A : Type} (K : Π{a : A} (p : a = a), p = idp) : is_set A :=
|
||||||
is_hset.mk _ (λa b p q, eq.rec_on q K p)
|
is_set.mk _ (λa b p q, eq.rec_on q K p)
|
||||||
|
|
||||||
theorem is_hset_of_relation.{u} {A : Type.{u}} (R : A → A → Type.{u})
|
theorem is_set_of_relation.{u} {A : Type.{u}} (R : A → A → Type.{u})
|
||||||
(mere : Π(a b : A), is_hprop (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) : is_hset A :=
|
(imp : Π{a b : A}, R a b → a = b) : is_set A :=
|
||||||
is_hset_of_axiom_K
|
is_set_of_axiom_K
|
||||||
(λa p,
|
(λa p,
|
||||||
have H2 : transport (λx, R a x → a = x) p (@imp a a) = @imp a a, from !apd,
|
have H2 : transport (λx, R a x → a = x) p (@imp a a) = @imp a a, from !apd,
|
||||||
have H3 : Π(r : R a a), transport (λx, a = x) p (imp r)
|
have H3 : Π(r : R a a), transport (λx, a = x) p (imp r)
|
||||||
|
@ -103,27 +103,27 @@ namespace is_trunc
|
||||||
calc
|
calc
|
||||||
imp (refl a) ⬝ p = transport (λx, a = x) p (imp (refl a)) : transport_eq_r
|
imp (refl a) ⬝ p = transport (λx, a = x) p (imp (refl a)) : transport_eq_r
|
||||||
... = imp (transport (λx, R a x) p (refl a)) : H3
|
... = imp (transport (λx, R a x) p (refl a)) : H3
|
||||||
... = imp (refl a) : is_hprop.elim,
|
... = imp (refl a) : is_prop.elim,
|
||||||
cancel_left (imp (refl a)) H4)
|
cancel_left (imp (refl a)) H4)
|
||||||
|
|
||||||
definition relation_equiv_eq {A : Type} (R : A → A → Type)
|
definition relation_equiv_eq {A : Type} (R : A → A → Type)
|
||||||
(mere : Π(a b : A), is_hprop (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 :=
|
||||||
@equiv_of_is_hprop _ _ _
|
@equiv_of_is_prop _ _ _
|
||||||
(@is_trunc_eq _ _ (is_hset_of_relation R mere refl @imp) a b)
|
(@is_trunc_eq _ _ (is_set_of_relation R mere refl @imp) a b)
|
||||||
imp
|
imp
|
||||||
(λp, p ▸ refl a)
|
(λp, p ▸ refl a)
|
||||||
|
|
||||||
local attribute not [reducible]
|
local attribute not [reducible]
|
||||||
theorem is_hset_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)
|
||||||
: is_hset A :=
|
: is_set A :=
|
||||||
is_hset_of_relation (λa b, ¬¬a = b) _ (λa n, n idp) H
|
is_set_of_relation (λa b, ¬¬a = b) _ (λa n, n idp) H
|
||||||
|
|
||||||
section
|
section
|
||||||
open decidable
|
open decidable
|
||||||
--this is proven differently in init.hedberg
|
--this is proven differently in init.hedberg
|
||||||
theorem is_hset_of_decidable_eq (A : Type) [H : decidable_eq A] : is_hset A :=
|
theorem is_set_of_decidable_eq (A : Type) [H : decidable_eq A] : is_set A :=
|
||||||
is_hset_of_double_neg_elim (λa b, by_contradiction)
|
is_set_of_double_neg_elim (λa b, by_contradiction)
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem is_trunc_of_axiom_K_of_leq {A : Type} (n : trunc_index) (H : -1 ≤ n)
|
theorem is_trunc_of_axiom_K_of_leq {A : Type} (n : trunc_index) (H : -1 ≤ n)
|
||||||
|
@ -138,9 +138,9 @@ namespace is_trunc
|
||||||
induction p, apply Hp
|
induction p, apply Hp
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem is_hprop_iff_is_contr {A : Type} (a : A) :
|
theorem is_prop_iff_is_contr {A : Type} (a : A) :
|
||||||
is_hprop A ↔ is_contr A :=
|
is_prop A ↔ is_contr A :=
|
||||||
iff.intro (λH, is_contr.mk a (is_hprop.elim a)) _
|
iff.intro (λH, is_contr.mk a (is_prop.elim a)) _
|
||||||
|
|
||||||
theorem is_trunc_succ_iff_is_trunc_loop (A : Type) (Hn : -1 ≤ n) :
|
theorem is_trunc_succ_iff_is_trunc_loop (A : Type) (Hn : -1 ≤ n) :
|
||||||
is_trunc (n.+1) A ↔ Π(a : A), is_trunc n (a = a) :=
|
is_trunc (n.+1) A ↔ Π(a : A), is_trunc n (a = a) :=
|
||||||
|
@ -152,7 +152,7 @@ namespace is_trunc
|
||||||
revert A, induction n with n IH,
|
revert A, induction n with n IH,
|
||||||
{ intro A, esimp [Iterated_loop_space], transitivity _,
|
{ intro A, esimp [Iterated_loop_space], transitivity _,
|
||||||
{ apply is_trunc_succ_iff_is_trunc_loop, apply le.refl},
|
{ apply is_trunc_succ_iff_is_trunc_loop, apply le.refl},
|
||||||
{ apply pi_iff_pi, intro a, esimp, apply is_hprop_iff_is_contr, reflexivity}},
|
{ apply pi_iff_pi, intro a, esimp, apply is_prop_iff_is_contr, reflexivity}},
|
||||||
{ intro A, esimp [Iterated_loop_space],
|
{ intro A, esimp [Iterated_loop_space],
|
||||||
transitivity _, apply @is_trunc_succ_iff_is_trunc_loop @n, esimp, constructor,
|
transitivity _, apply @is_trunc_succ_iff_is_trunc_loop @n, esimp, constructor,
|
||||||
apply pi_iff_pi, intro a, transitivity _, apply IH,
|
apply pi_iff_pi, intro a, transitivity _, apply IH,
|
||||||
|
@ -167,7 +167,7 @@ namespace is_trunc
|
||||||
induction n with n,
|
induction n with n,
|
||||||
{ esimp [sub_two,Iterated_loop_space], apply iff.intro,
|
{ esimp [sub_two,Iterated_loop_space], apply iff.intro,
|
||||||
intro H a, exact is_contr_of_inhabited_hprop a,
|
intro H a, exact is_contr_of_inhabited_hprop a,
|
||||||
intro H, apply is_hprop_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
|
||||||
|
|
||||||
|
@ -225,15 +225,15 @@ namespace trunc
|
||||||
{ 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_leq _ -2), exact unit.star},
|
{ apply (@is_trunc_of_leq _ -2), exact unit.star},
|
||||||
{ 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_hprop)),
|
apply (@trunc.rec_on _ _ _ aa (λy, !is_trunc_succ_of_is_prop)),
|
||||||
eapply (@trunc.rec_on _ _ _ aa' (λy, !is_trunc_succ_of_is_hprop)),
|
eapply (@trunc.rec_on _ _ _ aa' (λy, !is_trunc_succ_of_is_prop)),
|
||||||
intro a a', apply (is_trunc_equiv_closed_rev),
|
intro a a', apply (is_trunc_equiv_closed_rev),
|
||||||
{ apply tr_eq_tr_equiv},
|
{ apply tr_eq_tr_equiv},
|
||||||
{ exact (IH _ _ _)}}}
|
{ exact (IH _ _ _)}}}
|
||||||
end
|
end
|
||||||
|
|
||||||
open equiv.ops
|
open equiv.ops
|
||||||
definition unique_choice {P : A → Type} [H : Πa, is_hprop (P a)] (f : Πa, ∥ P a ∥) (a : A)
|
definition unique_choice {P : A → Type} [H : Πa, is_prop (P a)] (f : Πa, ∥ P a ∥) (a : A)
|
||||||
: P a :=
|
: P a :=
|
||||||
!trunc_equiv (f a)
|
!trunc_equiv (f a)
|
||||||
|
|
||||||
|
@ -262,7 +262,7 @@ 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_hprop (λ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))
|
||||||
|
|
||||||
end function
|
end function
|
||||||
|
|
|
@ -39,13 +39,13 @@ namespace univ
|
||||||
|
|
||||||
/- Properties which can be disproven for the universe -/
|
/- Properties which can be disproven for the universe -/
|
||||||
|
|
||||||
definition not_is_hset_type0 : ¬is_hset Type₀ :=
|
definition not_is_set_type0 : ¬is_set Type₀ :=
|
||||||
assume H : is_hset Type₀,
|
assume H : is_set Type₀,
|
||||||
absurd !is_hset.elim eq_bnot_ne_idp
|
absurd !is_set.elim eq_bnot_ne_idp
|
||||||
|
|
||||||
definition not_is_hset_type : ¬is_hset Type.{u} :=
|
definition not_is_set_type : ¬is_set Type.{u} :=
|
||||||
assume H : is_hset Type,
|
assume H : is_set Type,
|
||||||
absurd (is_trunc_is_embedding_closed lift star) not_is_hset_type0
|
absurd (is_trunc_is_embedding_closed lift star) not_is_set_type0
|
||||||
|
|
||||||
definition not_double_negation_elimination0 : ¬Π(A : Type₀), ¬¬A → A :=
|
definition not_double_negation_elimination0 : ¬Π(A : Type₀), ¬¬A → A :=
|
||||||
begin
|
begin
|
||||||
|
@ -53,7 +53,7 @@ namespace univ
|
||||||
have u : ¬¬bool, by exact (λg, g tt),
|
have u : ¬¬bool, by exact (λg, g tt),
|
||||||
let H1 := apdo f eq_bnot,
|
let H1 := apdo f eq_bnot,
|
||||||
note H2 := apo10 H1 u,
|
note H2 := apo10 H1 u,
|
||||||
have p : eq_bnot ▸ u = u, from !is_hprop.elim,
|
have p : eq_bnot ▸ u = u, from !is_prop.elim,
|
||||||
rewrite p at H2,
|
rewrite p at H2,
|
||||||
note H3 := eq_of_pathover_ua H2, esimp at H3, --TODO: use apply ... at after #700
|
note H3 := eq_of_pathover_ua H2, esimp at H3, --TODO: use apply ... at after #700
|
||||||
exact absurd H3 (bnot_ne (f bool u)),
|
exact absurd H3 (bnot_ne (f bool u)),
|
||||||
|
|
|
@ -189,10 +189,10 @@ optional<expr> mk_class_instance(environment const & env, old_local_context cons
|
||||||
return mk_class_instance(env, ctx.get_data(), type, nullptr);
|
return mk_class_instance(env, ctx.get_data(), type, nullptr);
|
||||||
}
|
}
|
||||||
|
|
||||||
optional<expr> mk_hset_instance(type_checker & tc, options const & o, list<expr> const & ctx, expr const & type) {
|
optional<expr> mk_set_instance(type_checker & tc, options const & o, list<expr> const & ctx, expr const & type) {
|
||||||
level lvl = sort_level(tc.ensure_type(type).first);
|
level lvl = sort_level(tc.ensure_type(type).first);
|
||||||
expr is_hset = tc.whnf(mk_app(mk_constant(get_is_trunc_is_hset_name(), {lvl}), type)).first;
|
expr is_set = tc.whnf(mk_app(mk_constant(get_is_trunc_is_set_name(), {lvl}), type)).first;
|
||||||
return mk_class_instance(tc.env(), o, ctx, is_hset);
|
return mk_class_instance(tc.env(), o, ctx, is_set);
|
||||||
}
|
}
|
||||||
|
|
||||||
optional<expr> mk_subsingleton_instance(environment const & env, options const & o, list<expr> const & ctx, expr const & type) {
|
optional<expr> mk_subsingleton_instance(environment const & env, options const & o, list<expr> const & ctx, expr const & type) {
|
||||||
|
|
|
@ -39,7 +39,7 @@ pair<expr, constraint> mk_class_instance_elaborator(
|
||||||
|
|
||||||
optional<expr> mk_class_instance(environment const & env, io_state const & ios, old_local_context const & ctx, expr const & type, bool use_local_instances);
|
optional<expr> mk_class_instance(environment const & env, io_state const & ios, old_local_context const & ctx, expr const & type, bool use_local_instances);
|
||||||
optional<expr> mk_class_instance(environment const & env, old_local_context const & ctx, expr const & type);
|
optional<expr> mk_class_instance(environment const & env, old_local_context const & ctx, expr const & type);
|
||||||
optional<expr> mk_hset_instance(type_checker & tc, options const & o, list<expr> const & ctx, expr const & type);
|
optional<expr> mk_set_instance(type_checker & tc, options const & o, list<expr> const & ctx, expr const & type);
|
||||||
optional<expr> mk_subsingleton_instance(environment const & env, options const & o,
|
optional<expr> mk_subsingleton_instance(environment const & env, options const & o,
|
||||||
list<expr> const & ctx, expr const & type);
|
list<expr> const & ctx, expr const & type);
|
||||||
|
|
||||||
|
|
|
@ -84,9 +84,9 @@ name const * g_implies = nullptr;
|
||||||
name const * g_implies_of_if_neg = nullptr;
|
name const * g_implies_of_if_neg = nullptr;
|
||||||
name const * g_implies_of_if_pos = nullptr;
|
name const * g_implies_of_if_pos = nullptr;
|
||||||
name const * g_implies_resolve = nullptr;
|
name const * g_implies_resolve = nullptr;
|
||||||
name const * g_is_trunc_is_hprop = nullptr;
|
name const * g_is_trunc_is_prop = nullptr;
|
||||||
name const * g_is_trunc_is_hprop_elim = nullptr;
|
name const * g_is_trunc_is_prop_elim = nullptr;
|
||||||
name const * g_is_trunc_is_hset = nullptr;
|
name const * g_is_trunc_is_set = nullptr;
|
||||||
name const * g_ite = nullptr;
|
name const * g_ite = nullptr;
|
||||||
name const * g_left_distrib = nullptr;
|
name const * g_left_distrib = nullptr;
|
||||||
name const * g_le_refl = nullptr;
|
name const * g_le_refl = nullptr;
|
||||||
|
@ -359,9 +359,9 @@ void initialize_constants() {
|
||||||
g_implies_of_if_neg = new name{"implies_of_if_neg"};
|
g_implies_of_if_neg = new name{"implies_of_if_neg"};
|
||||||
g_implies_of_if_pos = new name{"implies_of_if_pos"};
|
g_implies_of_if_pos = new name{"implies_of_if_pos"};
|
||||||
g_implies_resolve = new name{"implies", "resolve"};
|
g_implies_resolve = new name{"implies", "resolve"};
|
||||||
g_is_trunc_is_hprop = new name{"is_trunc", "is_hprop"};
|
g_is_trunc_is_prop = new name{"is_trunc", "is_prop"};
|
||||||
g_is_trunc_is_hprop_elim = new name{"is_trunc", "is_hprop", "elim"};
|
g_is_trunc_is_prop_elim = new name{"is_trunc", "is_prop", "elim"};
|
||||||
g_is_trunc_is_hset = new name{"is_trunc", "is_hset"};
|
g_is_trunc_is_set = new name{"is_trunc", "is_set"};
|
||||||
g_ite = new name{"ite"};
|
g_ite = new name{"ite"};
|
||||||
g_left_distrib = new name{"left_distrib"};
|
g_left_distrib = new name{"left_distrib"};
|
||||||
g_le_refl = new name{"le", "refl"};
|
g_le_refl = new name{"le", "refl"};
|
||||||
|
@ -635,9 +635,9 @@ void finalize_constants() {
|
||||||
delete g_implies_of_if_neg;
|
delete g_implies_of_if_neg;
|
||||||
delete g_implies_of_if_pos;
|
delete g_implies_of_if_pos;
|
||||||
delete g_implies_resolve;
|
delete g_implies_resolve;
|
||||||
delete g_is_trunc_is_hprop;
|
delete g_is_trunc_is_prop;
|
||||||
delete g_is_trunc_is_hprop_elim;
|
delete g_is_trunc_is_prop_elim;
|
||||||
delete g_is_trunc_is_hset;
|
delete g_is_trunc_is_set;
|
||||||
delete g_ite;
|
delete g_ite;
|
||||||
delete g_left_distrib;
|
delete g_left_distrib;
|
||||||
delete g_le_refl;
|
delete g_le_refl;
|
||||||
|
@ -910,9 +910,9 @@ name const & get_implies_name() { return *g_implies; }
|
||||||
name const & get_implies_of_if_neg_name() { return *g_implies_of_if_neg; }
|
name const & get_implies_of_if_neg_name() { return *g_implies_of_if_neg; }
|
||||||
name const & get_implies_of_if_pos_name() { return *g_implies_of_if_pos; }
|
name const & get_implies_of_if_pos_name() { return *g_implies_of_if_pos; }
|
||||||
name const & get_implies_resolve_name() { return *g_implies_resolve; }
|
name const & get_implies_resolve_name() { return *g_implies_resolve; }
|
||||||
name const & get_is_trunc_is_hprop_name() { return *g_is_trunc_is_hprop; }
|
name const & get_is_trunc_is_prop_name() { return *g_is_trunc_is_prop; }
|
||||||
name const & get_is_trunc_is_hprop_elim_name() { return *g_is_trunc_is_hprop_elim; }
|
name const & get_is_trunc_is_prop_elim_name() { return *g_is_trunc_is_prop_elim; }
|
||||||
name const & get_is_trunc_is_hset_name() { return *g_is_trunc_is_hset; }
|
name const & get_is_trunc_is_set_name() { return *g_is_trunc_is_set; }
|
||||||
name const & get_ite_name() { return *g_ite; }
|
name const & get_ite_name() { return *g_ite; }
|
||||||
name const & get_left_distrib_name() { return *g_left_distrib; }
|
name const & get_left_distrib_name() { return *g_left_distrib; }
|
||||||
name const & get_le_refl_name() { return *g_le_refl; }
|
name const & get_le_refl_name() { return *g_le_refl; }
|
||||||
|
|
|
@ -86,9 +86,9 @@ name const & get_implies_name();
|
||||||
name const & get_implies_of_if_neg_name();
|
name const & get_implies_of_if_neg_name();
|
||||||
name const & get_implies_of_if_pos_name();
|
name const & get_implies_of_if_pos_name();
|
||||||
name const & get_implies_resolve_name();
|
name const & get_implies_resolve_name();
|
||||||
name const & get_is_trunc_is_hprop_name();
|
name const & get_is_trunc_is_prop_name();
|
||||||
name const & get_is_trunc_is_hprop_elim_name();
|
name const & get_is_trunc_is_prop_elim_name();
|
||||||
name const & get_is_trunc_is_hset_name();
|
name const & get_is_trunc_is_set_name();
|
||||||
name const & get_ite_name();
|
name const & get_ite_name();
|
||||||
name const & get_left_distrib_name();
|
name const & get_left_distrib_name();
|
||||||
name const & get_le_refl_name();
|
name const & get_le_refl_name();
|
||||||
|
|
|
@ -79,9 +79,9 @@ implies
|
||||||
implies_of_if_neg
|
implies_of_if_neg
|
||||||
implies_of_if_pos
|
implies_of_if_pos
|
||||||
implies.resolve
|
implies.resolve
|
||||||
is_trunc.is_hprop
|
is_trunc.is_prop
|
||||||
is_trunc.is_hprop.elim
|
is_trunc.is_prop.elim
|
||||||
is_trunc.is_hset
|
is_trunc.is_set
|
||||||
ite
|
ite
|
||||||
left_distrib
|
left_distrib
|
||||||
le.refl
|
le.refl
|
||||||
|
|
|
@ -43,7 +43,7 @@ result::result(list<goal> const & gs, list<list<expr>> const & args, list<implem
|
||||||
|
|
||||||
The eq_rec_eq definition is of the form
|
The eq_rec_eq definition is of the form
|
||||||
|
|
||||||
definition eq_rec_eq.{l₁ l₂} {A : Type.{l₁}} {B : A → Type.{l₂}} [h : is_hset A] {a : A} (b : B a) (p : a = a) :
|
definition eq_rec_eq.{l₁ l₂} {A : Type.{l₁}} {B : A → Type.{l₂}} [h : is_set A] {a : A} (b : B a) (p : a = a) :
|
||||||
b = @eq.rec.{l₂ l₁} A a (λ (a' : A) (h : a = a'), B a') b a p :=
|
b = @eq.rec.{l₂ l₁} A a (λ (a' : A) (h : a = a'), B a') b a p :=
|
||||||
...
|
...
|
||||||
*/
|
*/
|
||||||
|
@ -56,8 +56,8 @@ optional<expr> apply_eq_rec_eq(type_checker & tc, io_state const & ios, list<exp
|
||||||
if (!is_local(p) || !is_eq_a_a(tc, mlocal_type(p)))
|
if (!is_local(p) || !is_eq_a_a(tc, mlocal_type(p)))
|
||||||
return none_expr();
|
return none_expr();
|
||||||
expr const & A = args[0];
|
expr const & A = args[0];
|
||||||
auto is_hset_A = mk_hset_instance(tc, ios.get_options(), ctx, A);
|
auto is_set_A = mk_set_instance(tc, ios.get_options(), ctx, A);
|
||||||
if (!is_hset_A)
|
if (!is_set_A)
|
||||||
return none_expr();
|
return none_expr();
|
||||||
levels ls = const_levels(eq_rec_fn);
|
levels ls = const_levels(eq_rec_fn);
|
||||||
level l2 = head(ls);
|
level l2 = head(ls);
|
||||||
|
@ -76,7 +76,7 @@ optional<expr> apply_eq_rec_eq(type_checker & tc, io_state const & ios, list<exp
|
||||||
expr a = args[1];
|
expr a = args[1];
|
||||||
expr b = args[3];
|
expr b = args[3];
|
||||||
expr r = mk_constant(get_eq_rec_eq_name(), {l1, l2});
|
expr r = mk_constant(get_eq_rec_eq_name(), {l1, l2});
|
||||||
return some_expr(mk_app({r, A, B, *is_hset_A, a, b, p}));
|
return some_expr(mk_app({r, A, B, *is_set_A, a, b, p}));
|
||||||
}
|
}
|
||||||
|
|
||||||
typedef inversion::implementation_ptr implementation_ptr;
|
typedef inversion::implementation_ptr implementation_ptr;
|
||||||
|
@ -540,7 +540,7 @@ class inversion_tac {
|
||||||
/** \brief Process goal of the form: Pi (H : eq.rec A s C a s p = b), R
|
/** \brief Process goal of the form: Pi (H : eq.rec A s C a s p = b), R
|
||||||
The idea is to reduce it to
|
The idea is to reduce it to
|
||||||
Pi (H : a = b), R
|
Pi (H : a = b), R
|
||||||
when A is a hset
|
when A is a set
|
||||||
and then invoke intro_next_eq recursively.
|
and then invoke intro_next_eq recursively.
|
||||||
|
|
||||||
\remark \c type is the type of \c g after some normalization
|
\remark \c type is the type of \c g after some normalization
|
||||||
|
|
|
@ -1874,7 +1874,7 @@ optional<expr> type_context::mk_subsingleton_instance(expr const & type) {
|
||||||
if (is_standard(m_env))
|
if (is_standard(m_env))
|
||||||
subsingleton = mk_app(mk_constant(get_subsingleton_name(), {lvl}), type);
|
subsingleton = mk_app(mk_constant(get_subsingleton_name(), {lvl}), type);
|
||||||
else
|
else
|
||||||
subsingleton = whnf(mk_app(mk_constant(get_is_trunc_is_hprop_name(), {lvl}), type));
|
subsingleton = whnf(mk_app(mk_constant(get_is_trunc_is_prop_name(), {lvl}), type));
|
||||||
auto r = mk_class_instance(subsingleton);
|
auto r = mk_class_instance(subsingleton);
|
||||||
m_ss_cache.insert(mk_pair(type, r));
|
m_ss_cache.insert(mk_pair(type, r));
|
||||||
return r;
|
return r;
|
||||||
|
|
|
@ -632,7 +632,7 @@ expr mk_subsingleton_elim(type_checker & tc, expr const & h, expr const & x, exp
|
||||||
if (is_standard(tc.env())) {
|
if (is_standard(tc.env())) {
|
||||||
r = mk_constant(get_subsingleton_elim_name(), {l});
|
r = mk_constant(get_subsingleton_elim_name(), {l});
|
||||||
} else {
|
} else {
|
||||||
r = mk_constant(get_is_trunc_is_hprop_elim_name(), {l});
|
r = mk_constant(get_is_trunc_is_prop_elim_name(), {l});
|
||||||
}
|
}
|
||||||
return mk_app({r, A, h, x, y});
|
return mk_app({r, A, h, x, y});
|
||||||
}
|
}
|
||||||
|
|
|
@ -170,7 +170,7 @@ expr mk_symm(type_checker & tc, expr const & H);
|
||||||
expr mk_trans(type_checker & tc, expr const & H1, expr const & H2);
|
expr mk_trans(type_checker & tc, expr const & H1, expr const & H2);
|
||||||
expr mk_subst(type_checker & tc, expr const & motive, expr const & x, expr const & y, expr const & xeqy, expr const & h);
|
expr mk_subst(type_checker & tc, expr const & motive, expr const & x, expr const & y, expr const & xeqy, expr const & h);
|
||||||
expr mk_subst(type_checker & tc, expr const & motive, expr const & xeqy, expr const & h);
|
expr mk_subst(type_checker & tc, expr const & motive, expr const & xeqy, expr const & h);
|
||||||
/** \brief Create an proof for x = y using subsingleton.elim (in standard mode) and is_trunc.is_hprop.elim (in HoTT mode) */
|
/** \brief Create an proof for x = y using subsingleton.elim (in standard mode) and is_trunc.is_prop.elim (in HoTT mode) */
|
||||||
expr mk_subsingleton_elim(type_checker & tc, expr const & h, expr const & x, expr const & y);
|
expr mk_subsingleton_elim(type_checker & tc, expr const & h, expr const & x, expr const & y);
|
||||||
|
|
||||||
/** \brief Return true iff \c e is a term of the form (eq.rec ....) */
|
/** \brief Return true iff \c e is a term of the form (eq.rec ....) */
|
||||||
|
|
|
@ -21,7 +21,7 @@ constants H₁ : a ≠ b
|
||||||
constants H₂ : a = c
|
constants H₂ : a = c
|
||||||
namespace tst attribute H₂ [simp] end tst
|
namespace tst attribute H₂ [simp] end tst
|
||||||
|
|
||||||
definition ne_is_hprop [instance] (a b : nat) : is_hprop (a ≠ b) :=
|
definition ne_is_prop [instance] (a b : nat) : is_prop (a ≠ b) :=
|
||||||
sorry
|
sorry
|
||||||
-- TODO(Daniel): simplifier seems to have applied unnecessary step (see: (eq.nrec ... (eq.refl ..)))
|
-- TODO(Daniel): simplifier seems to have applied unnecessary step (see: (eq.nrec ... (eq.refl ..)))
|
||||||
#simplify eq tst 0 (g a b H₁)
|
#simplify eq tst 0 (g a b H₁)
|
||||||
|
|
Loading…
Reference in a new issue