chore(hott): standardize names of homotopy_of_inv_homotopy_post and friends

This commit is contained in:
Ulrik Buchholtz 2016-03-14 19:11:21 -04:00 committed by Leonardo de Moura
parent 5257e282aa
commit 1c52062f1e
4 changed files with 74 additions and 67 deletions

View file

@ -225,11 +225,12 @@ namespace susp
begin begin
fconstructor, fconstructor,
{ intro a, induction a, { intro a, induction a,
{ reflexivity}, { reflexivity },
{ reflexivity}, { reflexivity },
{ apply eq_pathover, apply hdeg_square, { apply eq_pathover, apply hdeg_square,
rewrite [▸*,ap_compose' _ (psusp_functor f),↑psusp_functor,+elim_merid]}}, rewrite [▸*,ap_compose' _ (psusp_functor f),↑psusp_functor],
{ reflexivity} krewrite +susp.elim_merid } },
{ reflexivity }
end end
-- adjunction from Coq-HoTT -- adjunction from Coq-HoTT

View file

@ -198,10 +198,11 @@ namespace is_equiv
end end
section section
variables {A B C : Type} {f : A → B} [Hf : is_equiv f] {a : A} {b : B} {g : B → C} {h : A → C} variables {A B C : Type} {f : A → B} [Hf : is_equiv f]
include Hf include Hf
--Rewrite rules section rewrite_rules
variables {a : A} {b : B}
definition eq_of_eq_inv (p : a = f⁻¹ b) : f a = b := definition eq_of_eq_inv (p : a = f⁻¹ b) : f a = b :=
ap f p ⬝ right_inv f b ap f p ⬝ right_inv f b
@ -213,19 +214,41 @@ namespace is_equiv
definition eq_inv_of_eq (p : f a = b) : a = f⁻¹ b := definition eq_inv_of_eq (p : f a = b) : a = f⁻¹ b :=
(inv_eq_of_eq p⁻¹)⁻¹ (inv_eq_of_eq p⁻¹)⁻¹
end rewrite_rules
variable (f) variable (f)
definition homotopy_of_homotopy_inv' (p : g ~ h ∘ f⁻¹) : g ∘ f ~ h :=
λa, p (f a) ⬝ ap h (left_inv f a)
definition homotopy_of_inv_homotopy' (p : h ∘ f⁻¹ ~ g) : h ~ g ∘ f := section pre_compose
λa, ap h (left_inv f a)⁻¹ ⬝ p (f a) variables (α : A → C) (β : B → C)
definition inv_homotopy_of_homotopy' (p : h ~ g ∘ f) : h ∘ f⁻¹ ~ g := definition homotopy_of_homotopy_inv_pre (p : β ~ α ∘ f⁻¹) : β ∘ f ~ α :=
λb, p (f⁻¹ b) ⬝ ap g (right_inv f b) λ a, p (f a) ⬝ ap α (left_inv f a)
definition homotopy_inv_of_homotopy' (p : g ∘ f ~ h) : g ~ h ∘ f⁻¹ := definition homotopy_of_inv_homotopy_pre (p : α ∘ f⁻¹ ~ β) : α ~ β ∘ f :=
λb, ap g (right_inv f b)⁻¹ ⬝ p (f⁻¹ b) λ a, (ap α (left_inv f a))⁻¹ ⬝ p (f a)
definition inv_homotopy_of_homotopy_pre (p : α ~ β ∘ f) : α ∘ f⁻¹ ~ β :=
λ b, p (f⁻¹ b) ⬝ ap β (right_inv f b)
definition homotopy_inv_of_homotopy_pre (p : β ∘ f ~ α) : β ~ α ∘ f⁻¹ :=
λ b, (ap β (right_inv f b))⁻¹ ⬝ p (f⁻¹ b)
end pre_compose
section post_compose
variables (α : C → A) (β : C → B)
definition homotopy_of_homotopy_inv_post (p : α ~ f⁻¹ ∘ β) : f ∘ α ~ β :=
λ c, ap f (p c) ⬝ (right_inv f (β c))
definition homotopy_of_inv_homotopy_post (p : f⁻¹ ∘ β ~ α) : β ~ f ∘ α :=
λ c, (right_inv f (β c))⁻¹ ⬝ ap f (p c)
definition inv_homotopy_of_homotopy_post (p : β ~ f ∘ α) : f⁻¹ ∘ β ~ α :=
λ c, ap f⁻¹ (p c) ⬝ (left_inv f (α c))
definition homotopy_inv_of_homotopy_post (p : f ∘ α ~ β) : α ~ f⁻¹ ∘ β :=
λ c, (left_inv f (α c))⁻¹ ⬝ ap f⁻¹ (p c)
end post_compose
end end
@ -374,23 +397,6 @@ namespace equiv
end end
section
variables {A B C : Type} (f : A ≃ B) {a : A} {b : B} {g : B → C} {h : A → C}
definition homotopy_of_homotopy_inv (p : g ~ h ∘ f⁻¹) : g ∘ f ~ h :=
homotopy_of_homotopy_inv' f p
definition homotopy_of_inv_homotopy (p : h ∘ f⁻¹ ~ g) : h ~ g ∘ f :=
homotopy_of_inv_homotopy' f p
definition inv_homotopy_of_homotopy (p : h ~ g ∘ f) : h ∘ f⁻¹ ~ g :=
inv_homotopy_of_homotopy' f p
definition homotopy_inv_of_homotopy (p : g ∘ f ~ h) : g ~ h ∘ f⁻¹ :=
homotopy_inv_of_homotopy' f p
end
infixl ` ⬝pe `:75 := equiv_of_equiv_of_eq infixl ` ⬝pe `:75 := equiv_of_equiv_of_eq
infixl ` ⬝ep `:75 := equiv_of_eq_of_equiv infixl ` ⬝ep `:75 := equiv_of_eq_of_equiv

View file

@ -124,9 +124,9 @@ namespace arrow
open function open function
definition inv_commute_of_commute (p : f' ∘ α ~ β ∘ f) : f'⁻¹ ∘ β ~ α ∘ f⁻¹ := definition inv_commute_of_commute (p : f' ∘ α ~ β ∘ f) : f'⁻¹ ∘ β ~ α ∘ f⁻¹ :=
begin begin
apply homotopy_inv_of_homotopy_post f' β (α ∘ f⁻¹), apply inv_homotopy_of_homotopy_post f' (α ∘ f⁻¹) β,
apply homotopy.symm, apply homotopy.symm,
apply homotopy_inv_of_homotopy_pre f (f' ∘ α) β, apply inv_homotopy_of_homotopy_pre f (f' ∘ α) β,
apply p apply p
end end
@ -135,8 +135,8 @@ namespace arrow
= (ap f'⁻¹ (p a))⁻¹ ⬝ left_inv f' (α a) ⬝ ap α (left_inv f a)⁻¹ := = (ap f'⁻¹ (p a))⁻¹ ⬝ left_inv f' (α a) ⬝ ap α (left_inv f a)⁻¹ :=
begin begin
unfold inv_commute_of_commute, unfold inv_commute_of_commute,
unfold homotopy_inv_of_homotopy_post, unfold inv_homotopy_of_homotopy_post,
unfold homotopy_inv_of_homotopy_pre, unfold inv_homotopy_of_homotopy_pre,
rewrite [adj f a,-(ap_compose β f)], rewrite [adj f a,-(ap_compose β f)],
rewrite [eq_of_square (natural_square_tr p (left_inv f a))], rewrite [eq_of_square (natural_square_tr p (left_inv f a))],
rewrite [ap_inv f'⁻¹,ap_con f'⁻¹,con_inv,con.assoc,con.assoc], rewrite [ap_inv f'⁻¹,ap_con f'⁻¹,con_inv,con.assoc,con.assoc],
@ -152,8 +152,8 @@ namespace arrow
= right_inv f' (β b) ⬝ ap β (right_inv f b)⁻¹ ⬝ (p (f⁻¹ b))⁻¹ := = right_inv f' (β b) ⬝ ap β (right_inv f b)⁻¹ ⬝ (p (f⁻¹ b))⁻¹ :=
begin begin
unfold inv_commute_of_commute, unfold inv_commute_of_commute,
unfold homotopy_inv_of_homotopy_post, unfold inv_homotopy_of_homotopy_post,
unfold homotopy_inv_of_homotopy_pre, unfold inv_homotopy_of_homotopy_pre,
rewrite [ap_con,-(ap_compose f' f'⁻¹),-(adj f' (α (f⁻¹ b)))], rewrite [ap_con,-(ap_compose f' f'⁻¹),-(adj f' (α (f⁻¹ b)))],
rewrite [con.assoc (right_inv f' (β b)) (ap β (right_inv f b)⁻¹) rewrite [con.assoc (right_inv f' (β b)) (ap β (right_inv f b)⁻¹)
(p (f⁻¹ b))⁻¹], (p (f⁻¹ b))⁻¹],
@ -173,13 +173,13 @@ namespace arrow
begin begin
unfold inv_commute_of_commute, unfold inv_commute_of_commute,
apply @is_equiv_compose _ _ _ apply @is_equiv_compose _ _ _
(homotopy.symm ∘ (homotopy_inv_of_homotopy_pre f (f' ∘ α) β)) (homotopy.symm ∘ (inv_homotopy_of_homotopy_pre f (f' ∘ α) β))
(homotopy_inv_of_homotopy_post f' β (α ∘ f⁻¹)), (inv_homotopy_of_homotopy_post f' (α ∘ f⁻¹) β),
{ apply @is_equiv_compose _ _ _ { apply @is_equiv_compose _ _ _
(homotopy_inv_of_homotopy_pre f (f' ∘ α) β) homotopy.symm, (inv_homotopy_of_homotopy_pre f (f' ∘ α) β) homotopy.symm,
{ apply homotopy_inv_of_homotopy_pre.is_equiv }, { apply inv_homotopy_of_homotopy_pre.is_equiv },
{ apply pi.is_equiv_homotopy_symm } { apply pi.is_equiv_homotopy_symm }
}, },
{ apply homotopy_inv_of_homotopy_post.is_equiv } { apply inv_homotopy_of_homotopy_post.is_equiv }
end end
end arrow end arrow

View file

@ -111,15 +111,14 @@ namespace is_equiv
section pre_compose section pre_compose
variables (α : A → C) (β : B → C) variables (α : A → C) (β : B → C)
definition homotopy_inv_of_homotopy_pre (p : α ~ β ∘ f) : α ∘ f⁻¹ ~ β := -- homotopy_inv_of_homotopy_pre is in init.equiv
λb, p (f⁻¹ b) ⬝ ap β (right_inv f b) protected definition inv_homotopy_of_homotopy_pre.is_equiv
: is_equiv (inv_homotopy_of_homotopy_pre f α β) :=
protected definition homotopy_inv_of_homotopy_pre.is_equiv adjointify _ (homotopy_of_inv_homotopy_pre f α β)
: is_equiv (homotopy_inv_of_homotopy_pre f α β) :=
adjointify _ (λq a, (ap α (left_inv f a))⁻¹ ⬝ q (f a))
abstract begin abstract begin
intro q, apply eq_of_homotopy, intro b, intro q, apply eq_of_homotopy, intro b,
unfold homotopy_inv_of_homotopy_pre, unfold inv_homotopy_of_homotopy_pre,
unfold homotopy_of_inv_homotopy_pre,
apply inverse, apply eq_bot_of_square, apply inverse, apply eq_bot_of_square,
apply eq_hconcat (ap02 α (adj_inv f b)), apply eq_hconcat (ap02 α (adj_inv f b)),
apply eq_hconcat (ap_compose α f⁻¹ (right_inv f b))⁻¹, apply eq_hconcat (ap_compose α f⁻¹ (right_inv f b))⁻¹,
@ -127,7 +126,8 @@ namespace is_equiv
end end end end
abstract begin abstract begin
intro p, apply eq_of_homotopy, intro a, intro p, apply eq_of_homotopy, intro a,
unfold homotopy_inv_of_homotopy_pre, unfold inv_homotopy_of_homotopy_pre,
unfold homotopy_of_inv_homotopy_pre,
apply trans (con.assoc apply trans (con.assoc
(ap α (left_inv f a))⁻¹ (ap α (left_inv f a))⁻¹
(p (f⁻¹ (f a))) (p (f⁻¹ (f a)))
@ -140,17 +140,16 @@ namespace is_equiv
end pre_compose end pre_compose
section post_compose section post_compose
variables (β : C → B) (α : C → A) variables (α : C → A) (β : C → B)
definition homotopy_inv_of_homotopy_post (p : β ~ f ∘ α) : f⁻¹ ∘ β ~ α := -- homotopy_inv_of_homotopy_post is in init.equiv
λc, ap f⁻¹ (p c) ⬝ (left_inv f (α c)) protected definition inv_homotopy_of_homotopy_post.is_equiv
: is_equiv (inv_homotopy_of_homotopy_post f α β) :=
protected definition homotopy_inv_of_homotopy_post.is_equiv adjointify _ (homotopy_of_inv_homotopy_post f α β)
: is_equiv (homotopy_inv_of_homotopy_post f β α) :=
adjointify _ (λq c, (right_inv f (β c))⁻¹ ⬝ ap f (q c))
abstract begin abstract begin
intro q, apply eq_of_homotopy, intro c, intro q, apply eq_of_homotopy, intro c,
unfold homotopy_inv_of_homotopy_post, unfold inv_homotopy_of_homotopy_post,
unfold homotopy_of_inv_homotopy_post,
apply trans (whisker_right apply trans (whisker_right
(ap_con f⁻¹ (right_inv f (β c))⁻¹ (ap f (q c)) (ap_con f⁻¹ (right_inv f (β c))⁻¹ (ap f (q c))
⬝ whisker_right (ap_inv f⁻¹ (right_inv f (β c))) ⬝ whisker_right (ap_inv f⁻¹ (right_inv f (β c)))
@ -163,7 +162,8 @@ namespace is_equiv
end end end end
abstract begin abstract begin
intro p, apply eq_of_homotopy, intro c, intro p, apply eq_of_homotopy, intro c,
unfold homotopy_inv_of_homotopy_post, unfold inv_homotopy_of_homotopy_post,
unfold homotopy_of_inv_homotopy_post,
apply trans (whisker_left (right_inv f (β c))⁻¹ apply trans (whisker_left (right_inv f (β c))⁻¹
(ap_con f (ap f⁻¹ (p c)) (left_inv f (α c)))), (ap_con f (ap f⁻¹ (p c)) (left_inv f (α c)))),
apply trans (con.assoc (right_inv f (β c))⁻¹ (ap f (ap f⁻¹ (p c))) apply trans (con.assoc (right_inv f (β c))⁻¹ (ap f (ap f⁻¹ (p c)))