module Homotopy.Conjugation where
Conjugation of paths🔗
In any type for which we know two points , the existence of any identification induces an equivalence between the loop spaces and , given by transport in the usual way. However, since we know ahead-of-time what transport in a type of paths computes to, we can take a short-cut and define the equivalence directly: it is given by conjugation with .
opaque: ∀ {ℓ} {A : Type ℓ} {x y : A} → y ≡ x → y ≡ y → x ≡ x
conj = sym p ·· q ·· p conj p q
: (p : y ≡ x) (q : y ≡ y) → conj p q ≡ sym p ∙ q ∙ p
conj-defn = double-composite (sym p) q p
conj-defn p q
: (p : y ≡ x) (q : y ≡ y) → conj p q ≡ subst (λ x → x ≡ x) p q
conj-defn' = conj-defn p q ∙ sym (subst-path-both q p) conj-defn' p q
: (l : x ≡ x) → conj refl l ≡ l
conj-refl = conj-defn _ _ ·· ∙-idl _ ·· ∙-idr _
conj-refl l
: (p : x ≡ y) (q : y ≡ z) (r : x ≡ x)
conj-∙ → conj q (conj p r) ≡ conj (p ∙ q) r
= transport
conj-∙ p q r (λ i → conj-defn' q (conj-defn' p r (~ i)) (~ i) ≡ conj-defn' (p ∙ q) r (~ i))
(sym (subst-∙ (λ x → x ≡ x) p q r))
: (p : y ≡ x) → conj p refl ≡ refl
conj-of-refl = conj-defn _ _ ·· ap (sym p ∙_) (∙-idl p) ·· ∙-invl p
conj-of-refl p
: (p : y ≡ x) (q r : y ≡ y) → conj p (q ∙ r) ≡ conj p q ∙ conj p r
conj-of-∙ = J (λ x p → ∀ q r → conj p (q ∙ r) ≡ conj p q ∙ conj p r) λ q r →
conj-of-∙ (q ∙ r) ∙ ap₂ _∙_ (sym (conj-refl q)) (sym (conj-refl r)) conj-refl
opaque
unfolding conj
ap-conj: ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A}
→ (f : A → B) (p : y ≡ x) (q : y ≡ y)
→ ap f (conj p q) ≡ conj (ap f p) (ap f q)
= ap-·· f _ _ _ ap-conj f p q
opaque: conj (sym p) (conj p q) ≡ q
conj⁻conj {p = p} {q = q} =
conj⁻conj (conj _) (conj-defn' _ _)
ap _ _
·· conj-defn' (λ i → p i ≡ p i) q ·· transport⁻transport
opaque
pathp→conj: {p : y ≡ x} {q₁ : y ≡ y} {q₂ : x ≡ x}
→ PathP (λ i → p i ≡ p i) q₁ q₂ → conj p q₁ ≡ q₂
= conj-defn' _ _ ∙ from-pathp p pathp→conj p
opaque: {p q : x ≡ x} → q ∙ p ≡ p ∙ q → conj p q ≡ q
conj-commutative = conj-defn _ _ ·· ap₂ _∙_ refl α ·· ∙-cancell _ _ conj-commutative α
: (p : y ≡ x) → is-iso (conj p)
conj-is-iso .inv q = conj (sym p) q
conj-is-iso p .rinv q = conj⁻conj
conj-is-iso p .linv q = conj⁻conj conj-is-iso p