feat(hott/types): add characterization of lift, prove that Type.{u} is not an hset
This commit is contained in:
parent
9a439d4a4e
commit
cfddfdfa84
8 changed files with 254 additions and 59 deletions
|
@ -45,18 +45,19 @@ namespace is_equiv
|
|||
is_equiv.mk id id (λa, idp) (λa, idp) (λa, idp)
|
||||
|
||||
-- The composition of two equivalences is, again, an equivalence.
|
||||
definition is_equiv_compose [Hf : is_equiv f] [Hg : is_equiv g] : is_equiv (g ∘ f) :=
|
||||
is_equiv.mk (g ∘ f) (f⁻¹ ∘ g⁻¹)
|
||||
(λc, ap g (right_inv f (g⁻¹ c)) ⬝ right_inv g c)
|
||||
(λa, ap (inv f) (left_inv g (f a)) ⬝ left_inv f a)
|
||||
(λa, (whisker_left _ (adj g (f a))) ⬝
|
||||
(ap_con g _ _)⁻¹ ⬝
|
||||
ap02 g ((ap_con_eq_con (right_inv f) (left_inv g (f a)))⁻¹ ⬝
|
||||
(ap_compose f (inv f) _ ◾ adj f a) ⬝
|
||||
(ap_con f _ _)⁻¹
|
||||
) ⬝
|
||||
(ap_compose g f _)⁻¹
|
||||
)
|
||||
definition is_equiv_compose [constructor] [Hf : is_equiv f] [Hg : is_equiv g]
|
||||
: is_equiv (g ∘ f) :=
|
||||
is_equiv.mk (g ∘ f) (f⁻¹ ∘ g⁻¹)
|
||||
(λc, ap g (right_inv f (g⁻¹ c)) ⬝ right_inv g c)
|
||||
(λa, ap (inv f) (left_inv g (f a)) ⬝ left_inv f a)
|
||||
(λa, (whisker_left _ (adj g (f a))) ⬝
|
||||
(ap_con g _ _)⁻¹ ⬝
|
||||
ap02 g ((ap_con_eq_con (right_inv f) (left_inv g (f a)))⁻¹ ⬝
|
||||
(ap_compose f (inv f) _ ◾ adj f a) ⬝
|
||||
(ap_con f _ _)⁻¹
|
||||
) ⬝
|
||||
(ap_compose g f _)⁻¹
|
||||
)
|
||||
|
||||
-- Any function equal to an equivalence is an equivlance as well.
|
||||
variable {f}
|
||||
|
@ -106,7 +107,7 @@ namespace is_equiv
|
|||
end
|
||||
|
||||
-- Any function pointwise equal to an equivalence is an equivalence as well.
|
||||
definition homotopy_closed [constructor] {A B : Type} {f f' : A → B} [Hf : is_equiv f]
|
||||
definition homotopy_closed [constructor] {A B : Type} (f : A → B) {f' : A → B} [Hf : is_equiv f]
|
||||
(Hty : f ~ f') : is_equiv f' :=
|
||||
adjointify f'
|
||||
(inv f)
|
||||
|
@ -120,7 +121,7 @@ namespace is_equiv
|
|||
(λ b, ap f !Hty⁻¹ ⬝ right_inv f b)
|
||||
(λ a, !Hty⁻¹ ⬝ left_inv f a)
|
||||
|
||||
definition is_equiv_up [instance] (A : Type) : is_equiv (up : A → lift A) :=
|
||||
definition is_equiv_up [instance] [constructor] (A : Type) : is_equiv (up : A → lift A) :=
|
||||
adjointify up down (λa, by induction a;reflexivity) (λa, idp)
|
||||
|
||||
section
|
||||
|
@ -128,7 +129,7 @@ namespace is_equiv
|
|||
include Hf
|
||||
|
||||
--The inverse of an equivalence is, again, an equivalence.
|
||||
definition is_equiv_inv [instance] : is_equiv f⁻¹ :=
|
||||
definition is_equiv_inv [instance] [constructor] : is_equiv f⁻¹ :=
|
||||
adjointify f⁻¹ f (left_inv f) (right_inv f)
|
||||
|
||||
definition cancel_right (g : B → C) [Hgf : is_equiv (g ∘ f)] : (is_equiv g) :=
|
||||
|
@ -201,8 +202,9 @@ namespace is_equiv
|
|||
end
|
||||
|
||||
--Transporting is an equivalence
|
||||
definition is_equiv_tr [instance] {A : Type} (P : A → Type) {x y : A} (p : x = y) : (is_equiv (transport P p)) :=
|
||||
is_equiv.mk _ (transport P p⁻¹) (tr_inv_tr p) (inv_tr_tr p) (tr_inv_tr_lemma p)
|
||||
definition is_equiv_tr [instance] [constructor] {A : Type} (P : A → Type) {x y : A} (p : x = y)
|
||||
: (is_equiv (transport P p)) :=
|
||||
is_equiv.mk _ (transport P p⁻¹) (tr_inv_tr p) (inv_tr_tr p) (tr_inv_tr_lemma p)
|
||||
|
||||
end is_equiv
|
||||
open is_equiv
|
||||
|
@ -241,10 +243,10 @@ namespace equiv
|
|||
protected definition refl [refl] [constructor] : A ≃ A :=
|
||||
equiv.mk id !is_equiv_id
|
||||
|
||||
protected definition symm [symm] [unfold 3] (f : A ≃ B) : B ≃ A :=
|
||||
protected definition symm [symm] (f : A ≃ B) : B ≃ A :=
|
||||
equiv.mk f⁻¹ !is_equiv_inv
|
||||
|
||||
protected definition trans [trans] (f : A ≃ B) (g: B ≃ C) : A ≃ C :=
|
||||
protected definition trans [trans] (f : A ≃ B) (g : B ≃ C) : A ≃ C :=
|
||||
equiv.mk (g ∘ f) !is_equiv_compose
|
||||
|
||||
infixl `⬝e`:75 := equiv.trans
|
||||
|
@ -252,8 +254,12 @@ namespace equiv
|
|||
-- notation for inverse which is not overloaded
|
||||
abbreviation erfl [constructor] := @equiv.refl
|
||||
|
||||
definition to_inv_trans [reducible] [unfold-full] (f : A ≃ B) (g : B ≃ C)
|
||||
: to_inv (f ⬝e g) = to_fun (g⁻¹ᵉ ⬝e f⁻¹ᵉ) :=
|
||||
idp
|
||||
|
||||
definition equiv_change_fun [constructor] (f : A ≃ B) {f' : A → B} (Heq : f ~ f') : A ≃ B :=
|
||||
equiv.mk f' (is_equiv.homotopy_closed Heq)
|
||||
equiv.mk f' (is_equiv.homotopy_closed f Heq)
|
||||
|
||||
definition equiv_change_inv [constructor] (f : A ≃ B) {f' : B → A} (Heq : f⁻¹ ~ f')
|
||||
: A ≃ B :=
|
||||
|
|
|
@ -15,11 +15,16 @@ section
|
|||
universe variable l
|
||||
variables {A B : Type.{l}}
|
||||
|
||||
definition is_equiv_cast_of_eq (H : A = B) : is_equiv (cast H) :=
|
||||
(@is_equiv_tr Type (λX, X) A B H)
|
||||
definition is_equiv_cast_of_eq [constructor] (H : A = B) : is_equiv (cast H) :=
|
||||
is_equiv_tr (λX, X) H
|
||||
|
||||
definition equiv_of_eq [constructor] (H : A = B) : A ≃ B :=
|
||||
equiv.mk _ (is_equiv_cast_of_eq H)
|
||||
|
||||
definition equiv_of_eq_refl [reducible] [unfold-full] (A : Type)
|
||||
: equiv_of_eq (refl A) = equiv.refl :=
|
||||
idp
|
||||
|
||||
definition equiv_of_eq (H : A = B) : A ≃ B :=
|
||||
equiv.mk _ (is_equiv_cast_of_eq H)
|
||||
|
||||
end
|
||||
|
||||
|
@ -50,6 +55,9 @@ definition eq_of_equiv_lift {A B : Type} (f : A ≃ B) : A = lift B :=
|
|||
ua (f ⬝e !equiv_lift)
|
||||
|
||||
namespace equiv
|
||||
definition ua_refl (A : Type) : ua erfl = idpath A :=
|
||||
eq_of_fn_eq_fn !eq_equiv_equiv (right_inv !eq_equiv_equiv erfl)
|
||||
|
||||
-- One consequence of UA is that we can transport along equivalencies of types
|
||||
-- We can use this for calculation evironments
|
||||
protected definition transport_of_equiv [subst] (P : Type → Type) {A B : Type} (H : A ≃ B)
|
||||
|
|
|
@ -29,9 +29,6 @@ namespace bool
|
|||
assert H2 : bnot = id, from !cast_ua_fn⁻¹ ⬝ ap cast H,
|
||||
absurd (ap10 H2 tt) ff_ne_tt
|
||||
|
||||
definition not_is_hset_type : ¬is_hset Type₀ :=
|
||||
assume H : is_hset Type₀,
|
||||
absurd !is_hset.elim eq_bnot_ne_idp
|
||||
|
||||
definition bool_equiv_option_unit : bool ≃ option unit :=
|
||||
begin
|
||||
|
|
|
@ -109,7 +109,7 @@ end is_equiv
|
|||
|
||||
namespace equiv
|
||||
open is_equiv
|
||||
variables {A B : Type}
|
||||
variables {A B C : Type}
|
||||
|
||||
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' :=
|
||||
|
@ -118,6 +118,15 @@ namespace equiv
|
|||
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))
|
||||
|
||||
definition equiv_eq' {f f' : A ≃ B} (p : to_fun f ~ to_fun f') : f = f' :=
|
||||
by apply equiv_eq;apply eq_of_homotopy p
|
||||
|
||||
definition trans_symm (f : A ≃ B) (g : B ≃ C) : (f ⬝e g)⁻¹ᵉ = g⁻¹ᵉ ⬝e f⁻¹ᵉ :> (C ≃ A) :=
|
||||
equiv_eq idp
|
||||
|
||||
definition symm_symm (f : A ≃ B) : f⁻¹ᵉ⁻¹ᵉ = f :> (A ≃ B) :=
|
||||
equiv_eq idp
|
||||
|
||||
protected definition equiv.sigma_char [constructor]
|
||||
(A B : Type) : (A ≃ B) ≃ Σ(f : A → B), is_equiv f :=
|
||||
begin
|
||||
|
|
141
hott/types/lift.hlean
Normal file
141
hott/types/lift.hlean
Normal file
|
@ -0,0 +1,141 @@
|
|||
/-
|
||||
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Author: Floris van Doorn
|
||||
|
||||
Theorems about lift
|
||||
-/
|
||||
|
||||
import ..function
|
||||
open eq equiv equiv.ops is_equiv is_trunc
|
||||
|
||||
namespace lift
|
||||
|
||||
universe variables u v
|
||||
variables {A : Type.{u}} (z z' : lift.{u v} A)
|
||||
|
||||
protected definition eta : up (down z) = z :=
|
||||
by induction z; reflexivity
|
||||
|
||||
protected definition code [unfold 2 3] : lift A → lift A → Type
|
||||
| code (up a) (up a') := a = a'
|
||||
|
||||
protected definition decode [unfold 2 3] : Π(z z' : lift A), lift.code z z' → z = z'
|
||||
| decode (up a) (up a') := λc, ap up c
|
||||
|
||||
variables {z z'}
|
||||
protected definition encode [unfold 3 4 5] (p : z = z') : lift.code z z' :=
|
||||
by induction p; induction z; esimp
|
||||
|
||||
variables (z z')
|
||||
definition lift_eq_equiv : (z = z') ≃ lift.code z z' :=
|
||||
equiv.MK lift.encode
|
||||
!lift.decode
|
||||
abstract begin
|
||||
intro c, induction z with a, induction z' with a', esimp at *, induction c,
|
||||
reflexivity
|
||||
end end
|
||||
abstract begin
|
||||
intro p, induction p, induction z, reflexivity
|
||||
end end
|
||||
|
||||
|
||||
section
|
||||
variables {a a' : A}
|
||||
definition eq_of_up_eq_up [unfold 4] (p : up a = up a') : a = a' :=
|
||||
lift.encode p
|
||||
|
||||
definition lift_transport {P : A → Type} (p : a = a') (z : lift (P a))
|
||||
: p ▸ z = up (p ▸ down z) :=
|
||||
by induction p; induction z; reflexivity
|
||||
end
|
||||
|
||||
variables {A' : Type} (f : A → A') (g : lift A → lift A')
|
||||
definition lift_functor [unfold 4] : lift A → lift A'
|
||||
| lift_functor (up a) := up (f a)
|
||||
|
||||
definition is_equiv_lift_functor [constructor] [Hf : is_equiv f] : is_equiv (lift_functor f) :=
|
||||
adjointify (lift_functor f)
|
||||
(lift_functor f⁻¹)
|
||||
abstract begin
|
||||
intro z', induction z' with a',
|
||||
esimp, exact ap up !right_inv
|
||||
end end
|
||||
abstract begin
|
||||
intro z, induction z with a,
|
||||
esimp, exact ap up !left_inv
|
||||
end end
|
||||
|
||||
definition lift_equiv_lift_of_is_equiv [constructor] [Hf : is_equiv f] : lift A ≃ lift A' :=
|
||||
equiv.mk _ (is_equiv_lift_functor f)
|
||||
|
||||
definition lift_equiv_lift [constructor] (f : A ≃ A') : lift A ≃ lift A' :=
|
||||
equiv.mk _ (is_equiv_lift_functor f)
|
||||
|
||||
definition lift_equiv_lift_refl (A : Type) : lift_equiv_lift (erfl : A ≃ A) = erfl :=
|
||||
by apply equiv_eq'; intro z; induction z with a; reflexivity
|
||||
|
||||
definition lift_inv_functor [unfold-full] (a : A) : A' :=
|
||||
down (g (up a))
|
||||
|
||||
definition is_equiv_lift_inv_functor [constructor] [Hf : is_equiv g]
|
||||
: is_equiv (lift_inv_functor g) :=
|
||||
adjointify (lift_inv_functor g)
|
||||
(lift_inv_functor g⁻¹)
|
||||
abstract begin
|
||||
intro z', rewrite [▸*,lift.eta,right_inv g],
|
||||
end end
|
||||
abstract begin
|
||||
intro z', rewrite [▸*,lift.eta,left_inv g],
|
||||
end end
|
||||
|
||||
definition equiv_of_lift_equiv_lift [constructor] (g : lift A ≃ lift A') : A ≃ A' :=
|
||||
equiv.mk _ (is_equiv_lift_inv_functor g)
|
||||
|
||||
definition lift_functor_left_inv : lift_inv_functor (lift_functor f) = f :=
|
||||
eq_of_homotopy (λa, idp)
|
||||
|
||||
definition lift_functor_right_inv : lift_functor (lift_inv_functor g) = g :=
|
||||
begin
|
||||
apply eq_of_homotopy, intro z, induction z with a, esimp, apply lift.eta
|
||||
end
|
||||
|
||||
variables (A A')
|
||||
definition is_equiv_lift_functor_fn [constructor]
|
||||
: is_equiv (lift_functor : (A → A') → (lift A → lift A')) :=
|
||||
adjointify lift_functor
|
||||
lift_inv_functor
|
||||
lift_functor_right_inv
|
||||
lift_functor_left_inv
|
||||
|
||||
definition lift_imp_lift_equiv [constructor] : (lift A → lift A') ≃ (A → A') :=
|
||||
(equiv.mk _ (is_equiv_lift_functor_fn A A'))⁻¹ᵉ
|
||||
|
||||
-- can we deduce this from lift_imp_lift_equiv?
|
||||
definition lift_equiv_lift_equiv [constructor] : (lift A ≃ lift A') ≃ (A ≃ A') :=
|
||||
equiv.MK equiv_of_lift_equiv_lift
|
||||
lift_equiv_lift
|
||||
abstract begin
|
||||
intro f, apply equiv_eq, reflexivity
|
||||
end end
|
||||
abstract begin
|
||||
intro g, apply equiv_eq, esimp, apply eq_of_homotopy, intro z,
|
||||
induction z with a, esimp, apply lift.eta
|
||||
end end
|
||||
|
||||
definition lift_eq_lift_equiv.{u1 u2} (A A' : Type.{u1})
|
||||
: (lift.{u1 u2} A = lift.{u1 u2} A') ≃ (A = A') :=
|
||||
!eq_equiv_equiv ⬝e !lift_equiv_lift_equiv ⬝e !eq_equiv_equiv⁻¹ᵉ
|
||||
|
||||
definition is_embedding_lift [instance] : is_embedding lift :=
|
||||
begin
|
||||
apply is_embedding.mk, intro A A', fapply is_equiv.homotopy_closed,
|
||||
exact to_inv !lift_eq_lift_equiv,
|
||||
exact _,
|
||||
{ intro p, induction p,
|
||||
esimp [lift_eq_lift_equiv,equiv.trans,equiv.symm,eq_equiv_equiv],
|
||||
rewrite [equiv_of_eq_refl,lift_equiv_lift_refl],
|
||||
apply ua_refl}
|
||||
end
|
||||
|
||||
end lift
|
|
@ -68,14 +68,14 @@ namespace prod
|
|||
definition prod_eq_unc_eta (p : u = v) : prod_eq_unc (p..1, p..2) = p :=
|
||||
prod_eq_eta p
|
||||
|
||||
definition is_equiv_prod_eq [instance] (u v : A × B)
|
||||
definition is_equiv_prod_eq [instance] [constructor] (u v : A × B)
|
||||
: is_equiv (prod_eq_unc : u.1 = v.1 × u.2 = v.2 → u = v) :=
|
||||
adjointify prod_eq_unc
|
||||
(λp, (p..1, p..2))
|
||||
prod_eq_unc_eta
|
||||
pair_prod_eq_unc
|
||||
|
||||
definition prod_eq_equiv (u v : A × B) : (u = v) ≃ (u.1 = v.1 × u.2 = v.2) :=
|
||||
definition prod_eq_equiv [constructor] (u v : A × B) : (u = v) ≃ (u.1 = v.1 × u.2 = v.2) :=
|
||||
(equiv.mk prod_eq_unc _)⁻¹ᵉ
|
||||
|
||||
/- Transport -/
|
||||
|
@ -96,7 +96,7 @@ namespace prod
|
|||
|
||||
/- Equivalences -/
|
||||
|
||||
definition is_equiv_prod_functor [instance] [H : is_equiv f] [H : is_equiv g]
|
||||
definition is_equiv_prod_functor [instance] [constructor] [H : is_equiv f] [H : is_equiv g]
|
||||
: is_equiv (prod_functor f g) :=
|
||||
begin
|
||||
apply adjointify _ (prod_functor f⁻¹ g⁻¹),
|
||||
|
@ -104,33 +104,34 @@ namespace prod
|
|||
intro u, induction u, rewrite [▸*,left_inv f,left_inv g],
|
||||
end
|
||||
|
||||
definition prod_equiv_prod_of_is_equiv [H : is_equiv f] [H : is_equiv g]
|
||||
definition prod_equiv_prod_of_is_equiv [constructor] [H : is_equiv f] [H : is_equiv g]
|
||||
: A × B ≃ A' × B' :=
|
||||
equiv.mk (prod_functor f g) _
|
||||
|
||||
definition prod_equiv_prod (f : A ≃ A') (g : B ≃ B') : A × B ≃ A' × B' :=
|
||||
definition prod_equiv_prod [constructor] (f : A ≃ A') (g : B ≃ B') : A × B ≃ A' × B' :=
|
||||
equiv.mk (prod_functor f g) _
|
||||
|
||||
definition prod_equiv_prod_left (g : B ≃ B') : A × B ≃ A × B' :=
|
||||
definition prod_equiv_prod_left [constructor] (g : B ≃ B') : A × B ≃ A × B' :=
|
||||
prod_equiv_prod equiv.refl g
|
||||
|
||||
definition prod_equiv_prod_right (f : A ≃ A') : A × B ≃ A' × B :=
|
||||
definition prod_equiv_prod_right [constructor] (f : A ≃ A') : A × B ≃ A' × B :=
|
||||
prod_equiv_prod f equiv.refl
|
||||
|
||||
/- Symmetry -/
|
||||
|
||||
definition is_equiv_flip [instance] (A B : Type) : is_equiv (@flip A B) :=
|
||||
definition is_equiv_flip [instance] [constructor] (A B : Type)
|
||||
: is_equiv (flip : A × B → B × A) :=
|
||||
adjointify flip
|
||||
flip
|
||||
(λu, destruct u (λb a, idp))
|
||||
(λu, destruct u (λa b, idp))
|
||||
|
||||
definition prod_comm_equiv (A B : Type) : A × B ≃ B × A :=
|
||||
definition prod_comm_equiv [constructor] (A B : Type) : A × B ≃ B × A :=
|
||||
equiv.mk flip _
|
||||
|
||||
/- Associativity -/
|
||||
|
||||
definition prod_assoc_equiv (A B C : Type) : A × (B × C) ≃ (A × B) × C :=
|
||||
definition prod_assoc_equiv [constructor] (A B C : Type) : A × (B × C) ≃ (A × B) × C :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ intro z, induction z with a z, induction z with b c, exact (a, b, c)},
|
||||
|
@ -139,24 +140,24 @@ namespace prod
|
|||
{ intro z, induction z with a z, induction z with b c, reflexivity},
|
||||
end
|
||||
|
||||
definition prod_contr_equiv (A B : Type) [H : is_contr B] : A × B ≃ A :=
|
||||
definition prod_contr_equiv [constructor] (A B : Type) [H : is_contr B] : A × B ≃ A :=
|
||||
equiv.MK pr1
|
||||
(λx, (x, !center))
|
||||
(λx, idp)
|
||||
(λx, by cases x with a b; exact pair_eq idp !center_eq)
|
||||
|
||||
definition prod_unit_equiv (A : Type) : A × unit ≃ A :=
|
||||
definition prod_unit_equiv [constructor] (A : Type) : A × unit ≃ A :=
|
||||
!prod_contr_equiv
|
||||
|
||||
/- Universal mapping properties -/
|
||||
definition is_equiv_prod_rec [instance] (P : A × B → Type)
|
||||
definition is_equiv_prod_rec [instance] [constructor] (P : A × B → Type)
|
||||
: is_equiv (prod.rec : (Πa b, P (a, b)) → Πu, P u) :=
|
||||
adjointify _
|
||||
(λg a b, g (a, b))
|
||||
(λg, eq_of_homotopy (λu, by induction u;reflexivity))
|
||||
(λf, idp)
|
||||
|
||||
definition equiv_prod_rec (P : A × B → Type) : (Πa b, P (a, b)) ≃ (Πu, P u) :=
|
||||
definition equiv_prod_rec [constructor] (P : A × B → Type) : (Πa b, P (a, b)) ≃ (Πu, P u) :=
|
||||
equiv.mk prod.rec _
|
||||
|
||||
definition imp_imp_equiv_prod_imp (A B C : Type) : (A → B → C) ≃ (A × B → C) :=
|
||||
|
@ -166,17 +167,19 @@ namespace prod
|
|||
: P a × Q a :=
|
||||
(u.1 a, u.2 a)
|
||||
|
||||
definition is_equiv_prod_corec (P Q : A → Type)
|
||||
definition is_equiv_prod_corec [constructor] (P Q : A → Type)
|
||||
: is_equiv (prod_corec_unc : (Πa, P a) × (Πa, Q a) → Πa, P a × Q a) :=
|
||||
adjointify _
|
||||
(λg, (λa, (g a).1, λa, (g a).2))
|
||||
(by intro g; apply eq_of_homotopy; intro a; esimp; induction (g a); reflexivity)
|
||||
(by intro h; induction h with f g; reflexivity)
|
||||
|
||||
definition equiv_prod_corec (P Q : A → Type) : ((Πa, P a) × (Πa, Q a)) ≃ (Πa, P a × Q a) :=
|
||||
definition equiv_prod_corec [constructor] (P Q : A → Type)
|
||||
: ((Πa, P a) × (Πa, Q a)) ≃ (Πa, P a × Q a) :=
|
||||
equiv.mk _ !is_equiv_prod_corec
|
||||
|
||||
definition imp_prod_imp_equiv_imp_prod (A B C : Type) : (A → B) × (A → C) ≃ (A → (B × C)) :=
|
||||
definition imp_prod_imp_equiv_imp_prod [constructor] (A B C : Type)
|
||||
: (A → B) × (A → C) ≃ (A → (B × C)) :=
|
||||
!equiv_prod_corec
|
||||
|
||||
definition is_trunc_prod (A B : Type) (n : trunc_index) [HA : is_trunc n A] [HB : is_trunc n B]
|
||||
|
|
|
@ -31,7 +31,7 @@ namespace sum
|
|||
by induction p; induction z; all_goals exact up idp
|
||||
|
||||
variables (z z')
|
||||
definition sum_eq_equiv : (z = z') ≃ sum.code z z' :=
|
||||
definition sum_eq_equiv [constructor] : (z = z') ≃ sum.code z z' :=
|
||||
equiv.MK sum.encode
|
||||
!sum.decode
|
||||
abstract begin
|
||||
|
@ -63,7 +63,8 @@ namespace sum
|
|||
| sum_functor (inl a) := inl (f a)
|
||||
| sum_functor (inr b) := inr (g b)
|
||||
|
||||
definition is_equiv_sum_functor [Hf : is_equiv f] [Hg : is_equiv g] : is_equiv (sum_functor f g) :=
|
||||
definition is_equiv_sum_functor [constructor] [Hf : is_equiv f] [Hg : is_equiv g]
|
||||
: is_equiv (sum_functor f g) :=
|
||||
adjointify (sum_functor f g)
|
||||
(sum_functor f⁻¹ g⁻¹)
|
||||
abstract begin
|
||||
|
@ -75,23 +76,24 @@ namespace sum
|
|||
all_goals (esimp; (apply ap inl | apply ap inr); apply right_inv)
|
||||
end end
|
||||
|
||||
definition sum_equiv_sum_of_is_equiv [Hf : is_equiv f] [Hg : is_equiv g] : A + B ≃ A' + B' :=
|
||||
definition sum_equiv_sum_of_is_equiv [constructor] [Hf : is_equiv f] [Hg : is_equiv g]
|
||||
: A + B ≃ A' + B' :=
|
||||
equiv.mk _ (is_equiv_sum_functor f g)
|
||||
|
||||
definition sum_equiv_sum (f : A ≃ A') (g : B ≃ B') : A + B ≃ A' + B' :=
|
||||
definition sum_equiv_sum [constructor] (f : A ≃ A') (g : B ≃ B') : A + B ≃ A' + B' :=
|
||||
equiv.mk _ (is_equiv_sum_functor f g)
|
||||
|
||||
definition sum_equiv_sum_left (g : B ≃ B') : A + B ≃ A + B' :=
|
||||
definition sum_equiv_sum_left [constructor] (g : B ≃ B') : A + B ≃ A + B' :=
|
||||
sum_equiv_sum equiv.refl g
|
||||
|
||||
definition sum_equiv_sum_right (f : A ≃ A') : A + B ≃ A' + B :=
|
||||
definition sum_equiv_sum_right [constructor] (f : A ≃ A') : A + B ≃ A' + B :=
|
||||
sum_equiv_sum f equiv.refl
|
||||
|
||||
definition flip : A + B → B + A
|
||||
definition flip [unfold 3] : A + B → B + A
|
||||
| flip (inl a) := inr a
|
||||
| flip (inr b) := inl b
|
||||
|
||||
definition sum_comm_equiv (A B : Type) : A + B ≃ B + A :=
|
||||
definition sum_comm_equiv [constructor] (A B : Type) : A + B ≃ B + A :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
exact flip,
|
||||
|
@ -107,7 +109,7 @@ namespace sum
|
|||
-- all_goals try (repeat (apply inl | apply inr | assumption); now),
|
||||
-- end
|
||||
|
||||
definition sum_empty_equiv (A : Type) : A + empty ≃ A :=
|
||||
definition sum_empty_equiv [constructor] (A : Type) : A + empty ≃ A :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
intro z, induction z, assumption, contradiction,
|
||||
|
@ -119,7 +121,7 @@ namespace sum
|
|||
definition sum_rec_unc {P : A + B → Type} (fg : (Πa, P (inl a)) × (Πb, P (inr b))) : Πz, P z :=
|
||||
sum.rec fg.1 fg.2
|
||||
|
||||
definition is_equiv_sum_rec (P : A + B → Type)
|
||||
definition is_equiv_sum_rec [constructor] (P : A + B → Type)
|
||||
: is_equiv (sum_rec_unc : (Πa, P (inl a)) × (Πb, P (inr b)) → Πz, P z) :=
|
||||
begin
|
||||
apply adjointify sum_rec_unc (λf, (λa, f (inl a), λb, f (inr b))),
|
||||
|
@ -127,13 +129,15 @@ namespace sum
|
|||
intro h, induction h with f g, reflexivity
|
||||
end
|
||||
|
||||
definition equiv_sum_rec (P : A + B → Type) : (Πa, P (inl a)) × (Πb, P (inr b)) ≃ Πz, P z :=
|
||||
definition equiv_sum_rec [constructor] (P : A + B → Type)
|
||||
: (Πa, P (inl a)) × (Πb, P (inr b)) ≃ Πz, P z :=
|
||||
equiv.mk _ !is_equiv_sum_rec
|
||||
|
||||
definition imp_prod_imp_equiv_sum_imp (A B C : Type) : (A → C) × (B → C) ≃ (A + B → C) :=
|
||||
definition imp_prod_imp_equiv_sum_imp [constructor] (A B C : Type)
|
||||
: (A → C) × (B → C) ≃ (A + B → C) :=
|
||||
!equiv_sum_rec
|
||||
|
||||
definition is_trunc_sum (n : trunc_index) [HA : is_trunc (n.+2) A] [HB : is_trunc (n.+2) B]
|
||||
definition is_trunc_sum (n : trunc_index) [HA : is_trunc (n.+2) A] [HB : is_trunc (n.+2) B]
|
||||
: is_trunc (n.+2) (A + B) :=
|
||||
begin
|
||||
apply is_trunc_succ_intro, intro z z',
|
||||
|
@ -150,7 +154,8 @@ namespace sum
|
|||
definition sigma_bool_of_sum {A B : Type} (z : A + B) : Σ(b : bool), bool.rec A B b :=
|
||||
by induction z with a b; exact ⟨ff, a⟩; exact ⟨tt, b⟩
|
||||
|
||||
definition sum_equiv_sigma_bool (A B : Type) : A + B ≃ Σ(b : bool), bool.rec A B b :=
|
||||
definition sum_equiv_sigma_bool [constructor] (A B : Type)
|
||||
: A + B ≃ Σ(b : bool), bool.rec A B b :=
|
||||
equiv.MK sigma_bool_of_sum
|
||||
sum_of_sigma_bool
|
||||
begin intro v, induction v with b x, induction b, all_goals reflexivity end
|
||||
|
|
26
hott/types/univ.hlean
Normal file
26
hott/types/univ.hlean
Normal file
|
@ -0,0 +1,26 @@
|
|||
/-
|
||||
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Author: Floris van Doorn
|
||||
|
||||
Theorems about the universe
|
||||
-/
|
||||
|
||||
-- see also init.ua
|
||||
|
||||
import .bool .trunc .lift
|
||||
|
||||
open is_trunc bool lift unit
|
||||
|
||||
namespace univ
|
||||
|
||||
definition not_is_hset_type0 : ¬is_hset Type₀ :=
|
||||
assume H : is_hset Type₀,
|
||||
absurd !is_hset.elim eq_bnot_ne_idp
|
||||
|
||||
definition not_is_hset_type.{u} : ¬is_hset Type.{u} :=
|
||||
assume H : is_hset Type,
|
||||
absurd (is_trunc_is_embedding_closed lift star) not_is_hset_type0
|
||||
|
||||
|
||||
end univ
|
Loading…
Reference in a new issue