feat(hott/types): prove that 'is_equiv f' is an hprop
This commit is contained in:
parent
704f2b2697
commit
3d7656078d
11 changed files with 293 additions and 233 deletions
|
@ -7,7 +7,7 @@ Authors: Floris van Doorn
|
|||
-/
|
||||
|
||||
--note: modify definition in category.set
|
||||
import algebra.category.constructions .morphism
|
||||
import algebra.category.constructions .iso
|
||||
|
||||
open category eq category.ops functor prod.ops is_trunc
|
||||
|
||||
|
|
|
@ -143,6 +143,9 @@ end funext
|
|||
|
||||
open funext
|
||||
|
||||
definition eq_equiv_homotopy {A : Type} {P : A → Type} {f g : Π x, P x} : (f = g) ≃ (f ∼ g) :=
|
||||
equiv.mk apD10 _
|
||||
|
||||
definition eq_of_homotopy {A : Type} {P : A → Type} {f g : Π x, P x} : f ∼ g → f = g :=
|
||||
(@apD10 A P f g)⁻¹
|
||||
|
||||
|
|
|
@ -549,10 +549,10 @@ namespace eq
|
|||
-- Unwhiskering, a.k.a. cancelling
|
||||
|
||||
definition cancel_left {x y z : A} (p : x = y) (q r : y = z) : (p ⬝ q = p ⬝ r) → (q = r) :=
|
||||
eq.rec_on p (λq r s, !idp_con⁻¹ ⬝ s ⬝ !idp_con) q r
|
||||
λs, !inv_con_cancel_left⁻¹ ⬝ whisker_left p⁻¹ s ⬝ !inv_con_cancel_left
|
||||
|
||||
definition cancel_right {x y z : A} (p q : x = y) (r : y = z) : (p ⬝ r = q ⬝ r) → (p = q) :=
|
||||
eq.rec_on r (λs, !con_idp⁻¹ ⬝ s ⬝ !con_idp)
|
||||
λs, !con_inv_cancel_right⁻¹ ⬝ whisker_right s r⁻¹ ⬝ !con_inv_cancel_right
|
||||
|
||||
-- Whiskering and identity paths.
|
||||
|
||||
|
@ -574,11 +574,11 @@ namespace eq
|
|||
|
||||
definition con2_idp {p q : x = y} (h : p = q) :
|
||||
h ◾ idp = whisker_right h idp :> (p ⬝ idp = q ⬝ idp) :=
|
||||
eq.rec_on h idp
|
||||
idp
|
||||
|
||||
definition idp_con2 {p q : x = y} (h : p = q) :
|
||||
idp ◾ h = whisker_left idp h :> (idp ⬝ p = idp ⬝ q) :=
|
||||
eq.rec_on h idp
|
||||
idp
|
||||
|
||||
-- The interchange law for concatenation.
|
||||
definition con2_con_con2 {p p' p'' : x = y} {q q' q'' : y = z}
|
||||
|
@ -590,6 +590,8 @@ namespace eq
|
|||
(whisker_right a q) ⬝ (whisker_left p' b) = (whisker_left p b) ⬝ (whisker_right a q') :=
|
||||
eq.rec_on b (eq.rec_on a (idp_con _)⁻¹)
|
||||
|
||||
|
||||
|
||||
-- Structure corresponding to the coherence equations of a bicategory.
|
||||
|
||||
-- The "pentagonator": the 3-cell witnessing the associativity pentagon.
|
||||
|
|
49
hott/types/arrow.hlean
Normal file
49
hott/types/arrow.hlean
Normal file
|
@ -0,0 +1,49 @@
|
|||
/-
|
||||
Copyright (c) 2014 Floris van Doorn. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Module: types.arrow
|
||||
Author: Floris van Doorn
|
||||
|
||||
Ported from Coq HoTT
|
||||
Theorems about arrow types (function spaces)
|
||||
-/
|
||||
|
||||
import types.pi
|
||||
|
||||
open eq equiv is_equiv funext pi equiv.ops
|
||||
|
||||
namespace arrow
|
||||
|
||||
variables {A A' : Type} {B B' : Type} {C : A → B → Type}
|
||||
{a a' a'' : A} {b b' b'' : B} {f g : A → B}
|
||||
|
||||
-- all lemmas here are special cases of the ones for pi-types
|
||||
|
||||
/- Functorial action -/
|
||||
variables (f0 : A' → A) (f1 : B → B')
|
||||
|
||||
definition arrow_functor : (A → B) → (A' → B') := pi_functor f0 (λa, f1)
|
||||
|
||||
/- Equivalences -/
|
||||
|
||||
definition is_equiv_arrow_functor
|
||||
[H0 : is_equiv f0] [H1 : is_equiv f1] : is_equiv (arrow_functor f0 f1) :=
|
||||
is_equiv_pi_functor f0 (λa, f1)
|
||||
|
||||
definition arrow_equiv_arrow_rev (f0 : A' ≃ A) (f1 : B ≃ B') : (A → B) ≃ (A' → B') :=
|
||||
equiv.mk _ (is_equiv_arrow_functor f0 f1)
|
||||
|
||||
definition arrow_equiv_arrow (f0 : A ≃ A') (f1 : B ≃ B') : (A → B) ≃ (A' → B') :=
|
||||
arrow_equiv_arrow_rev (equiv.symm f0) f1
|
||||
|
||||
definition arrow_equiv_arrow_right (f1 : B ≃ B') : (A → B) ≃ (A → B') :=
|
||||
arrow_equiv_arrow_rev equiv.refl f1
|
||||
|
||||
definition arrow_equiv_arrow_left_rev (f0 : A' ≃ A) : (A → B) ≃ (A' → B) :=
|
||||
arrow_equiv_arrow_rev f0 equiv.refl
|
||||
|
||||
definition arrow_equiv_arrow_left (f0 : A ≃ A') : (A → B) ≃ (A' → B) :=
|
||||
arrow_equiv_arrow f0 equiv.refl
|
||||
|
||||
end arrow
|
|
@ -14,9 +14,43 @@ open eq sigma sigma.ops equiv is_equiv
|
|||
namespace eq
|
||||
/- Path spaces -/
|
||||
|
||||
variables {A B : Type} {a a1 a2 a3 a4 : A} {b b1 b2 : B} {f g : A → B} {h : B → A}
|
||||
|
||||
/- The path spaces of a path space are not, of course, determined; they are just the
|
||||
higher-dimensional structure of the original space. -/
|
||||
|
||||
/- some lemmas about whiskering -/
|
||||
|
||||
definition whisker_left_con_right (p : a1 = a2) {q q' q'' : a2 = a3} (r : q = q') (s : q' = q'')
|
||||
: whisker_left p (r ⬝ s) = whisker_left p r ⬝ whisker_left p s :=
|
||||
begin
|
||||
cases p, cases r, cases s, apply idp
|
||||
end
|
||||
|
||||
definition whisker_right_con_right {p p' p'' : a1 = a2} (q : a2 = a3) (r : p = p') (s : p' = p'')
|
||||
: whisker_right (r ⬝ s) q = whisker_right r q ⬝ whisker_right s q :=
|
||||
begin
|
||||
cases q, cases r, cases s, apply idp
|
||||
end
|
||||
|
||||
definition whisker_left_con_left (p : a1 = a2) (p' : a2 = a3) {q q' : a3 = a4} (r : q = q')
|
||||
: whisker_left (p ⬝ p') r = !con.assoc ⬝ whisker_left p (whisker_left p' r) ⬝ !con.assoc' :=
|
||||
begin
|
||||
cases p', cases p, cases r, cases q, apply idp
|
||||
end
|
||||
|
||||
definition whisker_right_con_left {p p' : a1 = a2} (q : a2 = a3) (q' : a3 = a4) (r : p = p')
|
||||
: whisker_right r (q ⬝ q') = !con.assoc' ⬝ whisker_right (whisker_right r q) q' ⬝ !con.assoc :=
|
||||
begin
|
||||
cases q', cases q, cases r, cases p, apply idp
|
||||
end
|
||||
|
||||
definition whisker_left_inv_left (p : a2 = a1) {q q' : a2 = a3} (r : q = q')
|
||||
: !con_inv_cancel_left⁻¹ ⬝ whisker_left p (whisker_left p⁻¹ r) ⬝ !con_inv_cancel_left = r :=
|
||||
begin
|
||||
cases p, cases r, cases q, apply idp
|
||||
end
|
||||
|
||||
/- Transporting in path spaces.
|
||||
|
||||
There are potentially a lot of these lemmas, so we adopt a uniform naming scheme:
|
||||
|
@ -26,17 +60,15 @@ namespace eq
|
|||
- `F` means application of a function to that (varying) endpoint.
|
||||
-/
|
||||
|
||||
variables {A B : Type} {a a1 a2 a3 a4 : A} {b b1 b2 : B} {f g : A → B} {h : B → A}
|
||||
|
||||
definition transport_paths_l (p : a1 = a2) (q : a1 = a3)
|
||||
definition transport_eq_l (p : a1 = a2) (q : a1 = a3)
|
||||
: transport (λx, x = a3) p q = p⁻¹ ⬝ q :=
|
||||
by cases p; cases q; apply idp
|
||||
|
||||
definition transport_paths_r (p : a2 = a3) (q : a1 = a2)
|
||||
definition transport_eq_r (p : a2 = a3) (q : a1 = a2)
|
||||
: transport (λx, a1 = x) p q = q ⬝ p :=
|
||||
by cases p; cases q; apply idp
|
||||
|
||||
definition transport_paths_lr (p : a1 = a2) (q : a1 = a1)
|
||||
definition transport_eq_lr (p : a1 = a2) (q : a1 = a1)
|
||||
: transport (λx, x = x) p q = p⁻¹ ⬝ q ⬝ p :=
|
||||
begin
|
||||
apply (eq.rec_on p),
|
||||
|
@ -45,15 +77,15 @@ namespace eq
|
|||
apply idp_con
|
||||
end
|
||||
|
||||
definition transport_paths_Fl (p : a1 = a2) (q : f a1 = b)
|
||||
definition transport_eq_Fl (p : a1 = a2) (q : f a1 = b)
|
||||
: transport (λx, f x = b) p q = (ap f p)⁻¹ ⬝ q :=
|
||||
by cases p; cases q; apply idp
|
||||
|
||||
definition transport_paths_Fr (p : a1 = a2) (q : b = f a1)
|
||||
definition transport_eq_Fr (p : a1 = a2) (q : b = f a1)
|
||||
: transport (λx, b = f x) p q = q ⬝ (ap f p) :=
|
||||
by cases p; apply idp
|
||||
|
||||
definition transport_paths_FlFr (p : a1 = a2) (q : f a1 = g a1)
|
||||
definition transport_eq_FlFr (p : a1 = a2) (q : f a1 = g a1)
|
||||
: transport (λx, f x = g x) p q = (ap f p)⁻¹ ⬝ q ⬝ (ap g p) :=
|
||||
begin
|
||||
apply (eq.rec_on p),
|
||||
|
@ -62,7 +94,7 @@ namespace eq
|
|||
apply idp_con
|
||||
end
|
||||
|
||||
definition transport_paths_FlFr_D {B : A → Type} {f g : Πa, B a}
|
||||
definition transport_eq_FlFr_D {B : A → Type} {f g : Πa, B a}
|
||||
(p : a1 = a2) (q : f a1 = g a1)
|
||||
: transport (λx, f x = g x) p q = (apD f p)⁻¹ ⬝ ap (transport B p) q ⬝ (apD g p) :=
|
||||
begin
|
||||
|
@ -73,7 +105,7 @@ namespace eq
|
|||
apply ap_id
|
||||
end
|
||||
|
||||
definition transport_paths_FFlr (p : a1 = a2) (q : h (f a1) = a1)
|
||||
definition transport_eq_FFlr (p : a1 = a2) (q : h (f a1) = a1)
|
||||
: transport (λx, h (f x) = x) p q = (ap h (ap f p))⁻¹ ⬝ q ⬝ p :=
|
||||
begin
|
||||
apply (eq.rec_on p),
|
||||
|
@ -82,7 +114,7 @@ namespace eq
|
|||
apply idp_con,
|
||||
end
|
||||
|
||||
definition transport_paths_lFFr (p : a1 = a2) (q : a1 = h (f a1))
|
||||
definition transport_eq_lFFr (p : a1 = a2) (q : a1 = h (f a1))
|
||||
: transport (λx, x = h (f x)) p q = p⁻¹ ⬝ q ⬝ (ap h (ap f p)) :=
|
||||
begin
|
||||
apply (eq.rec_on p),
|
||||
|
@ -104,174 +136,136 @@ namespace eq
|
|||
|
||||
/- Path operations are equivalences -/
|
||||
|
||||
definition isequiv_path_inverse [instance] (a1 a2 : A) : is_equiv (@inverse A a1 a2) :=
|
||||
definition is_equiv_eq_inverse (a1 a2 : A) : is_equiv (@inverse A a1 a2) :=
|
||||
is_equiv.mk inverse inv_inv inv_inv (λp, eq.rec_on p idp)
|
||||
local attribute is_equiv_eq_inverse [instance]
|
||||
|
||||
definition equiv_path_inverse (a1 a2 : A) : (a1 = a2) ≃ (a2 = a1) :=
|
||||
definition eq_equiv_eq_symm (a1 a2 : A) : (a1 = a2) ≃ (a2 = a1) :=
|
||||
equiv.mk inverse _
|
||||
|
||||
definition isequiv_concat_l [instance] (p : a1 = a2) (a3 : A)
|
||||
definition is_equiv_concat_left [instance] (p : a1 = a2) (a3 : A)
|
||||
: is_equiv (@concat _ a1 a2 a3 p) :=
|
||||
is_equiv.mk (concat p⁻¹)
|
||||
(con_inv_cancel_left p)
|
||||
(inv_con_cancel_left p)
|
||||
(eq.rec_on p (λq, eq.rec_on q idp))
|
||||
local attribute is_equiv_concat_left [instance]
|
||||
|
||||
definition equiv_concat_l (p : a1 = a2) (a3 : A) : (a1 = a3) ≃ (a2 = a3) :=
|
||||
definition equiv_eq_closed_left (p : a1 = a2) (a3 : A) : (a1 = a3) ≃ (a2 = a3) :=
|
||||
equiv.mk (concat p⁻¹) _
|
||||
|
||||
definition isequiv_concat_r [instance] (p : a2 = a3) (a1 : A)
|
||||
definition is_equiv_concat_right [instance] (p : a2 = a3) (a1 : A)
|
||||
: is_equiv (λq : a1 = a2, q ⬝ p) :=
|
||||
is_equiv.mk (λq, q ⬝ p⁻¹)
|
||||
(λq, inv_con_cancel_right q p)
|
||||
(λq, con_inv_cancel_right q p)
|
||||
(eq.rec_on p (λq, eq.rec_on q idp))
|
||||
local attribute is_equiv_concat_right [instance]
|
||||
|
||||
definition equiv_concat_r (p : a2 = a3) (a1 : A) : (a1 = a2) ≃ (a1 = a3) :=
|
||||
definition equiv_eq_closed_right (p : a2 = a3) (a1 : A) : (a1 = a2) ≃ (a1 = a3) :=
|
||||
equiv.mk (λq, q ⬝ p) _
|
||||
|
||||
definition equiv_concat_lr (p : a1 = a2) (q : a3 = a4) : (a1 = a3) ≃ (a2 = a4) :=
|
||||
equiv.trans (equiv_concat_l p a3) (equiv_concat_r q a2)
|
||||
definition eq_equiv_eq_closed (p : a1 = a2) (q : a3 = a4) : (a1 = a3) ≃ (a2 = a4) :=
|
||||
equiv.trans (equiv_eq_closed_left p a3) (equiv_eq_closed_right q a2)
|
||||
|
||||
definition is_equiv_whisker_left (p : a1 = a2) (q r : a2 = a3)
|
||||
: is_equiv (@whisker_left A a1 a2 a3 p q r) :=
|
||||
begin
|
||||
fapply adjointify,
|
||||
{intro s, apply (!cancel_left s)},
|
||||
{intro s,
|
||||
apply concat, {apply whisker_left_con_right},
|
||||
apply concat, rotate_left 1, apply (whisker_left_inv_left p s),
|
||||
apply concat2,
|
||||
{apply concat, {apply whisker_left_con_right},
|
||||
apply concat2,
|
||||
{cases p, cases q, apply idp},
|
||||
{apply idp}},
|
||||
{cases p, cases r, apply idp}},
|
||||
{intro s, cases s, cases q, cases p, apply idp}
|
||||
end
|
||||
|
||||
definition eq_equiv_con_eq_con_left (p : a1 = a2) (q r : a2 = a3) : (q = r) ≃ (p ⬝ q = p ⬝ r) :=
|
||||
equiv.mk _ !is_equiv_whisker_left
|
||||
|
||||
definition is_equiv_whisker_right {p q : a1 = a2} (r : a2 = a3)
|
||||
: is_equiv (λs, @whisker_right A a1 a2 a3 p q s r) :=
|
||||
begin
|
||||
fapply adjointify,
|
||||
{intro s, apply (!cancel_right s)},
|
||||
{intro s, cases r, cases s, cases q, apply idp},
|
||||
{intro s, cases s, cases r, cases p, apply idp}
|
||||
end
|
||||
|
||||
definition eq_equiv_con_eq_con_right (p q : a1 = a2) (r : a2 = a3) : (p = q) ≃ (p ⬝ r = q ⬝ r) :=
|
||||
equiv.mk _ !is_equiv_whisker_right
|
||||
|
||||
definition is_equiv_con_eq_of_eq_inv_con (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: is_equiv (con_eq_of_eq_inv_con p q r) :=
|
||||
begin
|
||||
cases r,
|
||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right),
|
||||
end
|
||||
|
||||
definition eq_inv_con_equiv_con_eq (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: (p = r⁻¹ ⬝ q) ≃ (r ⬝ p = q) :=
|
||||
equiv.mk _ !is_equiv_con_eq_of_eq_inv_con
|
||||
|
||||
definition is_equiv_con_eq_of_eq_con_inv (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: is_equiv (con_eq_of_eq_con_inv p q r) :=
|
||||
begin
|
||||
cases p,
|
||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||
end
|
||||
|
||||
definition eq_con_inv_equiv_con_eq (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: (r = q ⬝ p⁻¹) ≃ (r ⬝ p = q) :=
|
||||
equiv.mk _ !is_equiv_con_eq_of_eq_con_inv
|
||||
|
||||
definition is_equiv_inv_con_eq_of_eq_con (p : a1 = a3) (q : a2 = a3) (r : a1 = a2)
|
||||
: is_equiv (inv_con_eq_of_eq_con p q r) :=
|
||||
begin
|
||||
cases r,
|
||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||
end
|
||||
|
||||
definition eq_con_equiv_inv_con_eq (p : a1 = a3) (q : a2 = a3) (r : a1 = a2)
|
||||
: (p = r ⬝ q) ≃ (r⁻¹ ⬝ p = q) :=
|
||||
equiv.mk _ !is_equiv_inv_con_eq_of_eq_con
|
||||
|
||||
definition is_equiv_con_inv_eq_of_eq_con (p : a3 = a1) (q : a2 = a3) (r : a2 = a1)
|
||||
: is_equiv (con_inv_eq_of_eq_con p q r) :=
|
||||
begin
|
||||
cases p,
|
||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||
end
|
||||
|
||||
definition eq_con_equiv_con_inv_eq (p : a3 = a1) (q : a2 = a3) (r : a2 = a1)
|
||||
: (r = q ⬝ p) ≃ (r ⬝ p⁻¹ = q) :=
|
||||
equiv.mk _ !is_equiv_con_inv_eq_of_eq_con
|
||||
|
||||
definition is_equiv_eq_con_of_inv_con_eq (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: is_equiv (eq_con_of_inv_con_eq p q r) :=
|
||||
begin
|
||||
cases r,
|
||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||
end
|
||||
|
||||
definition inv_con_eq_equiv_eq_con (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: (r⁻¹ ⬝ q = p) ≃ (q = r ⬝ p) :=
|
||||
equiv.mk _ !is_equiv_eq_con_of_inv_con_eq
|
||||
|
||||
definition is_equiv_eq_con_of_con_inv_eq (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: is_equiv (eq_con_of_con_inv_eq p q r) :=
|
||||
begin
|
||||
cases p,
|
||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||
end
|
||||
|
||||
definition con_inv_eq_equiv_eq_con (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: (q ⬝ p⁻¹ = r) ≃ (q = r ⬝ p) :=
|
||||
equiv.mk _ !is_equiv_eq_con_of_con_inv_eq
|
||||
|
||||
-- a lot of this library still needs to be ported from Coq HoTT
|
||||
|
||||
-- set_option pp.beta true
|
||||
-- check @cancel_left
|
||||
-- set_option pp.full_names true
|
||||
-- definition isequiv_whiskerL [instance] (p : a1 = a2) (q r : a2 = a3)
|
||||
-- : is_equiv (@whisker_left A a1 a2 a3 p q r) :=
|
||||
-- begin
|
||||
-- fapply adjointify,
|
||||
-- {intro H, apply (!cancel_left H)},
|
||||
-- {intro s, }
|
||||
-- -- reverts (q,r,a), apply (eq.rec_on p), esimp {whisker_left,concat2, idp, cancel_left, eq.rec_on}, intros, esimp,
|
||||
-- end
|
||||
-- check @whisker_right_con_whisker_left
|
||||
-- end
|
||||
/-begin
|
||||
refine (isequiv_adjointify _ _ _ _).
|
||||
- apply cancelL.
|
||||
- intros k. unfold cancelL.
|
||||
rewrite !whiskerL_pp.
|
||||
refine ((_ @@ 1 @@ _) ⬝ whiskerL_pVL p k).
|
||||
+ destruct p, q; reflexivity.
|
||||
+ destruct p, r; reflexivity.
|
||||
- intros k. unfold cancelL.
|
||||
refine ((_ @@ 1 @@ _) ⬝ whiskerL_VpL p k).
|
||||
+ destruct p, q; reflexivity.
|
||||
+ destruct p, r; reflexivity.
|
||||
end-/
|
||||
|
||||
|
||||
definition is_equiv_con_eq_of_eq_inv_con (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: is_equiv (con_eq_of_eq_inv_con p q r) :=
|
||||
begin
|
||||
cases r,
|
||||
apply (is_equiv_compose (λx, idp_con _ ⬝ x) (λx, x ⬝ idp_con _)),
|
||||
end
|
||||
|
||||
|
||||
definition equiv_moveR_Mp (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: (p = r⁻¹ ⬝ q) ≃ (r ⬝ p = q) :=
|
||||
calc
|
||||
(p = r⁻¹ ⬝ q) ≃ (r ⬝ p = r ⬝ (r⁻¹ ⬝ q)) : equiv_concat_l r
|
||||
... ≃ (r ⬝ p = q) : sorry
|
||||
|
||||
definition equiv_moveR_Mp (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||
: (p = r⁻¹ ⬝ q) ≃ (r ⬝ p = q) :=
|
||||
equiv. _ _ (con_eq_of_eq_inv_con p q r) _.
|
||||
|
||||
Global Instance isequiv_moveR_pM
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : y = x)
|
||||
: is_equiv (con_eq_of_eq_con_inv p q r).
|
||||
/-begin
|
||||
destruct p.
|
||||
apply (isequiv_compose' _ (isequiv_concat_l _ _) _ (isequiv_concat_r _ _)).
|
||||
end-/
|
||||
|
||||
definition equiv_moveR_pM
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : y = x)
|
||||
: (r = q ⬝ p⁻¹) ≃ (r ⬝ p = q) :=
|
||||
equiv.mk _ _ (con_eq_of_eq_con_inv p q r) _.
|
||||
|
||||
Global Instance isequiv_moveR_Vp
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : x = y)
|
||||
: is_equiv (inv_con_eq_of_eq_con p q r).
|
||||
/-begin
|
||||
destruct r.
|
||||
apply (isequiv_compose' _ (isequiv_concat_l _ _) _ (isequiv_concat_r _ _)).
|
||||
end-/
|
||||
|
||||
definition equiv_moveR_Vp
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : x = y)
|
||||
: (p = r ⬝ q) ≃ (r⁻¹ ⬝ p = q) :=
|
||||
equiv.mk _ _ (inv_con_eq_of_eq_con p q r) _.
|
||||
|
||||
Global Instance isequiv_moveR_pV
|
||||
{A : Type} {x y z : A} (p : z = x) (q : y = z) (r : y = x)
|
||||
: is_equiv (con_inv_eq_of_eq_con p q r).
|
||||
/-begin
|
||||
destruct p.
|
||||
apply (isequiv_compose' _ (isequiv_concat_l _ _) _ (isequiv_concat_r _ _)).
|
||||
end-/
|
||||
|
||||
definition equiv_moveR_pV
|
||||
{A : Type} {x y z : A} (p : z = x) (q : y = z) (r : y = x)
|
||||
: (r = q ⬝ p) ≃ (r ⬝ p⁻¹ = q) :=
|
||||
equiv.mk _ _ (con_inv_eq_of_eq_con p q r) _.
|
||||
|
||||
Global Instance isequiv_moveL_Mp
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : y = x)
|
||||
: is_equiv (eq_con_of_inv_con_eq p q r).
|
||||
/-begin
|
||||
destruct r.
|
||||
apply (isequiv_compose' _ (isequiv_concat_l _ _) _ (isequiv_concat_r _ _)).
|
||||
end-/
|
||||
|
||||
definition equiv_moveL_Mp
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : y = x)
|
||||
: (r⁻¹ ⬝ q = p) ≃ (q = r ⬝ p) :=
|
||||
equiv.mk _ _ (eq_con_of_inv_con_eq p q r) _.
|
||||
|
||||
definition isequiv_moveL_pM
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : y = x)
|
||||
: is_equiv (eq_con_of_con_inv_eq p q r).
|
||||
/-begin
|
||||
destruct p.
|
||||
apply (isequiv_compose' _ (isequiv_concat_l _ _) _ (isequiv_concat_r _ _)).
|
||||
end-/
|
||||
|
||||
definition equiv_moveL_pM
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : y = x) :
|
||||
q ⬝ p⁻¹ = r ≃ q = r ⬝ p :=
|
||||
equiv.mk _ _ _ (isequiv_moveL_pM p q r).
|
||||
|
||||
Global Instance isequiv_moveL_Vp
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : x = y)
|
||||
: is_equiv (eq_inv_con_of_con_eq p q r).
|
||||
/-begin
|
||||
destruct r.
|
||||
apply (isequiv_compose' _ (isequiv_concat_l _ _) _ (isequiv_concat_r _ _)).
|
||||
end-/
|
||||
|
||||
definition equiv_moveL_Vp
|
||||
{A : Type} {x y z : A} (p : x = z) (q : y = z) (r : x = y)
|
||||
: r ⬝ q = p ≃ q = r⁻¹ ⬝ p :=
|
||||
equiv.mk _ _ (eq_inv_con_of_con_eq p q r) _.
|
||||
|
||||
Global Instance isequiv_moveL_pV
|
||||
{A : Type} {x y z : A} (p : z = x) (q : y = z) (r : y = x)
|
||||
: is_equiv (eq_con_inv_of_con_eq p q r).
|
||||
/-begin
|
||||
destruct p.
|
||||
apply (isequiv_compose' _ (isequiv_concat_l _ _) _ (isequiv_concat_r _ _)).
|
||||
end-/
|
||||
|
||||
definition equiv_moveL_pV
|
||||
{A : Type} {x y z : A} (p : z = x) (q : y = z) (r : y = x)
|
||||
: q ⬝ p = r ≃ q = r ⬝ p⁻¹ :=
|
||||
equiv.mk _ _ (eq_con_inv_of_con_eq p q r) _.
|
||||
|
||||
|
||||
end eq
|
||||
|
|
|
@ -9,22 +9,79 @@ Ported from Coq HoTT
|
|||
Theorems about the types equiv and is_equiv
|
||||
-/
|
||||
|
||||
--import types.sigma
|
||||
import types.fiber types.arrow
|
||||
|
||||
open eq is_trunc
|
||||
open eq is_trunc sigma sigma.ops arrow pi
|
||||
|
||||
namespace is_equiv
|
||||
open equiv
|
||||
context
|
||||
parameters {A B : Type} (f : A → B) [H : is_equiv f]
|
||||
open equiv function
|
||||
section
|
||||
open fiber
|
||||
variables {A B : Type} (f : A → B) [H : is_equiv f]
|
||||
include H
|
||||
definition is_contr_fiber_of_is_equiv (b : B) : is_contr (fiber f b) :=
|
||||
is_contr.mk
|
||||
(fiber.mk (f⁻¹ b) (retr f b))
|
||||
(λz, fiber.rec_on z (λa p, fiber.eq_mk ((ap f⁻¹ p)⁻¹ ⬝ sect f a) (calc
|
||||
retr f b = (ap (f ∘ f⁻¹) p)⁻¹ ⬝ ((ap (f ∘ f⁻¹) p) ⬝ retr f b) : inv_con_cancel_left
|
||||
... = (ap (f ∘ f⁻¹) p)⁻¹ ⬝ (retr f (f a) ⬝ p) : by rewrite ap_con_eq_con
|
||||
... = (ap (f ∘ f⁻¹) p)⁻¹ ⬝ (ap f (sect f a) ⬝ p) : by rewrite adj
|
||||
... = (ap (f ∘ f⁻¹) p)⁻¹ ⬝ ap f (sect f a) ⬝ p : con.assoc
|
||||
... = (ap f (ap f⁻¹ p))⁻¹ ⬝ ap f (sect f a) ⬝ p : by rewrite ap_compose
|
||||
... = ap f (ap f⁻¹ p)⁻¹ ⬝ ap f (sect f a) ⬝ p : by rewrite ap_inv
|
||||
... = ap f ((ap f⁻¹ p)⁻¹ ⬝ sect f a) ⬝ p : by rewrite ap_con)))
|
||||
|
||||
definition is_contr_right_inverse : is_contr (Σ(g : B → A), f ∘ g ∼ id) :=
|
||||
begin
|
||||
fapply is_trunc_equiv_closed,
|
||||
{apply sigma_equiv_sigma_id, intro g, apply eq_equiv_homotopy},
|
||||
fapply is_trunc_equiv_closed,
|
||||
{apply fiber.sigma_char},
|
||||
fapply is_contr_fiber_of_is_equiv,
|
||||
apply (to_is_equiv (arrow_equiv_arrow_right (equiv.mk f H))),
|
||||
end
|
||||
|
||||
definition is_contr_right_coherence (u : Σ(g : B → A), f ∘ g ∼ id)
|
||||
: is_contr (Σ(η : u.1 ∘ f ∼ id), Π(a : A), u.2 (f a) = ap f (η a)) :=
|
||||
begin
|
||||
fapply is_trunc_equiv_closed,
|
||||
{apply equiv.symm, apply sigma_pi_equiv_pi_sigma},
|
||||
fapply is_trunc_equiv_closed,
|
||||
{apply pi_equiv_pi_id, intro a,
|
||||
apply (equiv_fiber_eq (fiber.mk (u.1 (f a)) (u.2 (f a))) (fiber.mk a idp))},
|
||||
fapply is_trunc_pi,
|
||||
intro a, fapply @is_contr_eq,
|
||||
apply is_contr_fiber_of_is_equiv
|
||||
end
|
||||
|
||||
end
|
||||
variables {A B : Type} (f : A → B)
|
||||
|
||||
protected definition sigma_char : (is_equiv f) ≃
|
||||
(Σ(g : B → A) (ε : f ∘ g ∼ id) (η : g ∘ f ∼ id), Π(a : A), ε (f a) = ap f (η a)) :=
|
||||
equiv.MK (λH, ⟨inv f, retr f, sect f, adj f⟩)
|
||||
(λp, is_equiv.mk p.1 p.2.1 p.2.2.1 p.2.2.2)
|
||||
(λp, begin
|
||||
cases p with (p1, p2),
|
||||
cases p2 with (p21, p22),
|
||||
cases p22 with (p221, p222),
|
||||
apply idp
|
||||
end)
|
||||
(λH, is_equiv.rec_on H (λH1 H2 H3 H4, idp))
|
||||
|
||||
protected definition sigma_char' : (is_equiv f) ≃
|
||||
(Σ(u : Σ(g : B → A), f ∘ g ∼ id), Σ(η : u.1 ∘ f ∼ id), Π(a : A), u.2 (f a) = ap f (η a)) :=
|
||||
calc
|
||||
(is_equiv f) ≃
|
||||
(Σ(g : B → A) (ε : f ∘ g ∼ id) (η : g ∘ f ∼ id), Π(a : A), ε (f a) = ap f (η a))
|
||||
: is_equiv.sigma_char
|
||||
... ≃ (Σ(u : Σ(g : B → A), f ∘ g ∼ id), Σ(η : u.1 ∘ f ∼ id), Π(a : A), u.2 (f a) = ap f (η a))
|
||||
: {sigma_assoc_equiv (λu, Σ(η : u.1 ∘ f ∼ id), Π(a : A), u.2 (f a) = ap f (η a))}
|
||||
|
||||
local attribute is_contr_right_inverse [instance]
|
||||
local attribute is_contr_right_coherence [instance]
|
||||
theorem is_hprop_is_equiv [instance] : is_hprop (is_equiv f) :=
|
||||
sorry
|
||||
is_hprop_of_imp_is_contr (λ(H : is_equiv f), is_trunc_equiv_closed -2 (equiv.symm !sigma_char'))
|
||||
|
||||
end is_equiv
|
||||
|
||||
|
|
|
@ -18,9 +18,9 @@ structure fiber {A B : Type} (f : A → B) (b : B) :=
|
|||
open equiv sigma sigma.ops eq
|
||||
|
||||
namespace fiber
|
||||
variables {A B : Type} (f : A → B) (b : B)
|
||||
variables {A B : Type} {f : A → B} {b : B}
|
||||
|
||||
definition sigma_char : fiber f b ≃ (Σ(a : A), f a = b) :=
|
||||
definition sigma_char (f : A → B) (b : B) : fiber f b ≃ (Σ(a : A), f a = b) :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{intro x, exact ⟨point x, point_eq x⟩},
|
||||
|
@ -41,44 +41,13 @@ namespace fiber
|
|||
apply sigma_equiv_sigma_id,
|
||||
intro p,
|
||||
apply equiv_of_equiv_of_eq,
|
||||
{apply (ap (λx, x = _)), apply transport_paths_Fl}
|
||||
|
||||
-- apply equiv_of_eq,
|
||||
-- fapply (apD011 @sigma),
|
||||
-- {apply idp},
|
||||
-- esimp
|
||||
{apply (ap (λx, x = _)), apply transport_eq_Fl},
|
||||
apply inv_con_eq_equiv_eq_con,
|
||||
end
|
||||
|
||||
definition fiber_eq : (p : u.1 = v.1) (q : p ▹ u.2 = v.2) : u = v :=
|
||||
by cases u; cases v; apply (dpair_eq_dpair p q)
|
||||
definition eq_mk {x y : fiber f b} (p : point x = point y) (q : point_eq x = ap f p ⬝ point_eq y)
|
||||
: x = y :=
|
||||
to_inv !equiv_fiber_eq ⟨p, q⟩
|
||||
|
||||
|
||||
end fiber
|
||||
|
||||
|
||||
|
||||
namespace is_equiv
|
||||
open equiv
|
||||
context
|
||||
parameters {A B : Type} (f : A → B) [H : is_equiv f]
|
||||
include H
|
||||
|
||||
|
||||
end
|
||||
variables {A B : Type} (f : A → B)
|
||||
|
||||
theorem is_hprop_is_equiv [instance] : is_hprop (is_equiv f) :=
|
||||
sorry
|
||||
|
||||
end is_equiv
|
||||
|
||||
namespace equiv
|
||||
open is_equiv
|
||||
variables {A B : Type}
|
||||
|
||||
protected definition eq_mk' {f f' : A → B} [H : is_equiv f] [H' : is_equiv f'] (p : f = f')
|
||||
: equiv.mk f H = equiv.mk f' H' :=
|
||||
apD011 equiv.mk p !is_hprop.elim
|
||||
|
||||
protected definition eq_mk {f f' : A ≃ B} (p : to_fun f = to_fun f') : f = f' :=
|
||||
by (cases f; cases f'; apply (equiv.eq_mk' p))
|
||||
end equiv
|
||||
|
|
|
@ -107,16 +107,16 @@ namespace pi
|
|||
definition ap_pi_functor {g g' : Π(a:A), B a} (h : g ∼ g')
|
||||
: ap (pi_functor f0 f1) (eq_of_homotopy h) = eq_of_homotopy (λa':A', (ap (f1 a') (h (f0 a')))) :=
|
||||
begin
|
||||
apply (equiv_rect (@apD10 A B g g')), intro p, clear h, --revert p, revert g',
|
||||
apply (eq.rec_on p),
|
||||
apply concat, --(@concat _ _ (refl (pi_functor f0 f1 g))),
|
||||
apply (equiv_rect (@apD10 A B g g')), intro p, clear h,
|
||||
cases p,
|
||||
apply concat,
|
||||
exact (ap (ap (pi_functor f0 f1)) (eq_of_homotopy_idp g)),
|
||||
apply symm, apply eq_of_homotopy_idp
|
||||
end
|
||||
|
||||
/- Equivalences -/
|
||||
|
||||
definition is_equiv_pi_functor [instance] (f0 : A' → A) (f1 : Π(a':A'), B (f0 a') → B' a')
|
||||
definition is_equiv_pi_functor [instance]
|
||||
[H0 : is_equiv f0] [H1 : Πa', @is_equiv (B (f0 a')) (B' a') (f1 a')]
|
||||
: is_equiv (pi_functor f0 f1) :=
|
||||
begin
|
||||
|
@ -144,15 +144,12 @@ namespace pi
|
|||
: (Πa, B a) ≃ (Πa', B' a') :=
|
||||
equiv.mk (pi_functor f0 f1) _
|
||||
|
||||
context
|
||||
attribute inv [irreducible] --this is needed for the following class instance resolution
|
||||
definition pi_equiv_pi (f0 : A' ≃ A) (f1 : Πa', (B (to_fun f0 a') ≃ B' a'))
|
||||
: (Πa, B a) ≃ (Πa', B' a') :=
|
||||
pi_equiv_pi_of_is_equiv (to_fun f0) (λa', to_fun (f1 a'))
|
||||
end
|
||||
|
||||
definition pi_equiv_pi_id {P Q : A → Type} (g : Πa, P a ≃ Q a) : (Πa, P a) ≃ (Πa, Q a) :=
|
||||
pi_equiv_pi equiv.refl g.
|
||||
pi_equiv_pi equiv.refl g
|
||||
|
||||
/- Truncatedness: any dependent product of n-types is an n-type -/
|
||||
|
||||
|
|
|
@ -277,22 +277,11 @@ namespace sigma
|
|||
|
||||
--this proof is harder than in Coq because we don't have eta definitionally for sigma
|
||||
definition sigma_assoc_equiv (C : (Σa, B a) → Type) : (Σa b, C ⟨a, b⟩) ≃ (Σu, C u) :=
|
||||
-- begin
|
||||
-- apply equiv.mk,
|
||||
-- apply (adjointify (λav, ⟨⟨av.1, av.2.1⟩, av.2.2⟩)
|
||||
-- (λuc, ⟨uc.1.1, uc.1.2, !eta⁻¹ ▹ uc.2⟩)),
|
||||
-- intro uc, apply (destruct uc), intro u,
|
||||
-- apply (destruct u), intros (a, b, c),
|
||||
-- apply idp,
|
||||
-- intro av, apply (destruct av), intros (a, v),
|
||||
-- apply (destruct v), intros (b, c),
|
||||
-- apply idp,
|
||||
-- end
|
||||
equiv.mk _ (adjointify
|
||||
(λav, ⟨⟨av.1, av.2.1⟩, av.2.2⟩)
|
||||
(λuc, ⟨uc.1.1, uc.1.2, !eta⁻¹ ▹ uc.2⟩)
|
||||
proof (λuc, destruct uc (λu, destruct u (λa b c, idp))) qed
|
||||
proof (λav, destruct av (λa v, destruct v (λb c, idp))) qed)
|
||||
begin intro uc; cases uc with (u, c); cases u; apply idp end
|
||||
begin intro av; cases av with (a, v); cases v; apply idp end)
|
||||
|
||||
open prod
|
||||
definition assoc_equiv_prod (C : (A × A') → Type) : (Σa a', C (a,a')) ≃ (Σu, C u) :=
|
||||
|
@ -358,7 +347,7 @@ namespace sigma
|
|||
(λ h, proof eq_of_homotopy (λu, !eta) qed)
|
||||
(λfg, destruct fg (λ(f : Π (a : A), B a) (g : Π (x : A), C x (f x)), proof idp qed))
|
||||
|
||||
definition equiv_coind : (Σ(f : Πa, B a), Πa, C a (f a)) ≃ (Πa, Σb, C a b) :=
|
||||
definition sigma_pi_equiv_pi_sigma : (Σ(f : Πa, B a), Πa, C a (f a)) ≃ (Πa, Σb, C a b) :=
|
||||
equiv.mk coind_uncurried _
|
||||
end
|
||||
|
||||
|
|
|
@ -8,9 +8,9 @@ Authors: Jakob von Raumer, Floris van Doorn
|
|||
Properties of is_trunc
|
||||
-/
|
||||
|
||||
import types.pi types.path
|
||||
import types.pi types.eq
|
||||
|
||||
open sigma sigma.ops pi function eq equiv path funext
|
||||
open sigma sigma.ops pi function eq equiv eq funext
|
||||
namespace is_trunc
|
||||
|
||||
definition is_contr.sigma_char (A : Type) :
|
||||
|
@ -86,7 +86,7 @@ namespace is_trunc
|
|||
to_fun (equiv.symm !heq_pi) H2,
|
||||
have H4 : imp (refl a) ⬝ p = imp (refl a), from
|
||||
calc
|
||||
imp (refl a) ⬝ p = transport (λx, a = x) p (imp (refl a)) : transport_paths_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 (refl a) : is_hprop.elim,
|
||||
cancel_left (imp (refl a)) _ _ H4)
|
||||
|
|
|
@ -643,7 +643,7 @@ order for the change to take effect."
|
|||
("Gb" . ("β")) ("GB" . ("Β"))
|
||||
("Gg" . ("γ")) ("GG" . ("Γ"))
|
||||
("Gd" . ("δ")) ("GD" . ("Δ"))
|
||||
("Ge" . ("ε")) ("GE" . ("Ε"))
|
||||
("Ge" . ("ε")) ("GE" . ("Ε")) ("eps" . ("ε"))
|
||||
("Gz" . ("ζ")) ("GZ" . ("Ζ"))
|
||||
;; \eta \Eta
|
||||
("Gth" . ("θ")) ("GTH" . ("Θ"))
|
||||
|
|
Loading…
Reference in a new issue