[Agda-dev] One more point for my false-golfing score

Nils Anders Danielsson nad at cse.gu.se
Wed Aug 19 15:14:50 CEST 2015


On 2015-08-18 22:12, Andrea Vezzosi wrote:
> Do you still have the version without abstract? I'd like to try with
> my lazy-unfolding patch

I've attached a patch.

-- 
/NAD
-------------- next part --------------
1 patch for repository http://www.cse.chalmers.se/~nad/repos/equality:

patch 510360bf5698fd3ac4eb55b907f5cc21f47e414d
Author: Nils Anders Danielsson <nils.anders.danielsson at gmail.com>
Date:   Tue Aug 18 19:13:19 CEST 2015
  * Removed every occurrence of the abstract keyword.

New patches:

[Removed every occurrence of the abstract keyword.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150818171319
 Ignore-this: 302fa056db4f4cedc8239ae2ed9c9156
] hunk ./Bag-equivalence.agda 469
-abstract
-
-  -- In fact, they have equal lengths.
-
-  length-cong : ∀ {a} {A : Set a} {xs ys : List A} →
-                xs ≈-bag ys → length xs ≡ length ys
-  length-cong = _⇔_.to Fin.isomorphic-same-size ∘ Fin-length-cong
-
-  -- All that remains (except for some bookkeeping) is to show that
-  -- the isomorphism which Fin-length-cong returns relates the two
-  -- lists.
-
-  Fin-length-cong-relates :
-    ∀ {a} {A : Set a} {xs ys : List A} (xs≈ys : xs ≈-bag ys) →
-    xs And ys Are-related-by Fin-length-cong xs≈ys
-  Fin-length-cong-relates {xs = xs} {ys} xs≈ys i =
-    lookup xs i                                 ≡⟨ proj₂ $ to (∈-lookup _) $ to (xs≈ys _) (from (∈-lookup _) (i , refl)) ⟩
-    lookup ys (proj₁ $ to (∈-lookup _) $
-               to (xs≈ys _) $
-               from (∈-lookup _) (i , refl))    ≡⟨ refl ⟩∎
-    lookup ys (to (Fin-length-cong xs≈ys) i)    ∎
-    where open _↔_
+-- In fact, they have equal lengths.
+
+length-cong : ∀ {a} {A : Set a} {xs ys : List A} →
+              xs ≈-bag ys → length xs ≡ length ys
+length-cong = _⇔_.to Fin.isomorphic-same-size ∘ Fin-length-cong
+
+-- All that remains (except for some bookkeeping) is to show that
+-- the isomorphism which Fin-length-cong returns relates the two
+-- lists.
+
+Fin-length-cong-relates :
+  ∀ {a} {A : Set a} {xs ys : List A} (xs≈ys : xs ≈-bag ys) →
+  xs And ys Are-related-by Fin-length-cong xs≈ys
+Fin-length-cong-relates {xs = xs} {ys} xs≈ys i =
+  lookup xs i                                 ≡⟨ proj₂ $ to (∈-lookup _) $ to (xs≈ys _) (from (∈-lookup _) (i , refl)) ⟩
+  lookup ys (proj₁ $ to (∈-lookup _) $
+             to (xs≈ys _) $
+             from (∈-lookup _) (i , refl))    ≡⟨ refl ⟩∎
+  lookup ys (to (Fin-length-cong xs≈ys) i)    ∎
+  where open _↔_
hunk ./Bag-equivalence.agda 534
-abstract
-
-  -- The index function commutes with applications of certain
-  -- inverses. Note that the last three equational reasoning steps do
-  -- not need to be written out; I included them in an attempt to make
-  -- it easier to understand why the lemma holds.
-
-  index-commutes : ∀ {a} {A : Set a} {z : A} {xs ys} →
-                   (xs≈ys : xs ≈-bag ys) (p : z ∈ xs) →
-                   index (_↔_.to (xs≈ys z) p) ≡
-                   _↔_.to (Fin-length-cong xs≈ys) (index p)
-  index-commutes {z = z} {xs} {ys} xs≈ys p =
-    (index $ to (xs≈ys z) p)                             ≡⟨ lemma z p ⟩
-    (index $ to (xs≈ys _) $ proj₂ $
-     from (Fin-length xs) $ to (Fin-length xs) (z , p))  ≡⟨ refl ⟩
-    (index $ proj₂ $ Σ-map P.id (to (xs≈ys _)) $
-     from (Fin-length xs) $ to (Fin-length xs) (z , p))  ≡⟨ refl ⟩
-    (to (Fin-length ys) $ Σ-map P.id (to (xs≈ys _)) $
-     from (Fin-length xs) $ index p)                     ≡⟨ refl ⟩∎
-    (to (Fin-length-cong xs≈ys) $ index p)               ∎
-    where
-    open _↔_
-
-    lemma : ∀ z p →
-            (index $ to (xs≈ys z) p) ≡
-            (index $
-             to (xs≈ys (lookup xs (to (Fin-length xs) (z , p)))) $
-             proj₂ $ from (Fin-length xs) $
-             to (Fin-length xs) (z , p))
-    lemma z p with to (Fin-length xs) (z , p)
-                 | Σ-≡,≡←≡ (left-inverse-of (Fin-length xs) (z , p))
-    lemma .(lookup xs i) .(from (∈-lookup xs) (i , refl))
-          | i | refl , refl = refl
-
-  -- Bag equivalence isomorphisms preserve index equality. Note that
-  -- this means that, even if the underlying equality is proof
-  -- relevant, a bag equivalence isomorphism cannot map two distinct
-  -- proofs of z ∈ xs (say) to different positions.
-
-  index-equality-preserved :
-    ∀ {a} {A : Set a} {z : A} {xs ys} {p q : z ∈ xs}
-    (xs≈ys : xs ≈-bag ys) →
-    index p ≡ index q →
-    index (_↔_.to (xs≈ys z) p) ≡ index (_↔_.to (xs≈ys z) q)
-  index-equality-preserved {z = z} {p = p} {q} xs≈ys eq =
-    index (_↔_.to (xs≈ys z) p)                ≡⟨ index-commutes xs≈ys p ⟩
-    _↔_.to (Fin-length-cong xs≈ys) (index p)  ≡⟨ cong (_↔_.to (Fin-length-cong xs≈ys)) eq ⟩
-    _↔_.to (Fin-length-cong xs≈ys) (index q)  ≡⟨ sym $ index-commutes xs≈ys q ⟩∎
-    index (_↔_.to (xs≈ys z) q)                ∎
+-- The index function commutes with applications of certain
+-- inverses. Note that the last three equational reasoning steps do
+-- not need to be written out; I included them in an attempt to make
+-- it easier to understand why the lemma holds.
+
+index-commutes : ∀ {a} {A : Set a} {z : A} {xs ys} →
+                 (xs≈ys : xs ≈-bag ys) (p : z ∈ xs) →
+                 index (_↔_.to (xs≈ys z) p) ≡
+                 _↔_.to (Fin-length-cong xs≈ys) (index p)
+index-commutes {z = z} {xs} {ys} xs≈ys p =
+  (index $ to (xs≈ys z) p)                             ≡⟨ lemma z p ⟩
+  (index $ to (xs≈ys _) $ proj₂ $
+   from (Fin-length xs) $ to (Fin-length xs) (z , p))  ≡⟨ refl ⟩
+  (index $ proj₂ $ Σ-map P.id (to (xs≈ys _)) $
+   from (Fin-length xs) $ to (Fin-length xs) (z , p))  ≡⟨ refl ⟩
+  (to (Fin-length ys) $ Σ-map P.id (to (xs≈ys _)) $
+   from (Fin-length xs) $ index p)                     ≡⟨ refl ⟩∎
+  (to (Fin-length-cong xs≈ys) $ index p)               ∎
+  where
+  open _↔_
+
+  lemma : ∀ z p →
+          (index $ to (xs≈ys z) p) ≡
+          (index $
+           to (xs≈ys (lookup xs (to (Fin-length xs) (z , p)))) $
+           proj₂ $ from (Fin-length xs) $
+           to (Fin-length xs) (z , p))
+  lemma z p with to (Fin-length xs) (z , p)
+               | Σ-≡,≡←≡ (left-inverse-of (Fin-length xs) (z , p))
+  lemma .(lookup xs i) .(from (∈-lookup xs) (i , refl))
+        | i | refl , refl = refl
+
+-- Bag equivalence isomorphisms preserve index equality. Note that
+-- this means that, even if the underlying equality is proof
+-- relevant, a bag equivalence isomorphism cannot map two distinct
+-- proofs of z ∈ xs (say) to different positions.
+
+index-equality-preserved :
+  ∀ {a} {A : Set a} {z : A} {xs ys} {p q : z ∈ xs}
+  (xs≈ys : xs ≈-bag ys) →
+  index p ≡ index q →
+  index (_↔_.to (xs≈ys z) p) ≡ index (_↔_.to (xs≈ys z) q)
+index-equality-preserved {z = z} {p = p} {q} xs≈ys eq =
+  index (_↔_.to (xs≈ys z) p)                ≡⟨ index-commutes xs≈ys p ⟩
+  _↔_.to (Fin-length-cong xs≈ys) (index p)  ≡⟨ cong (_↔_.to (Fin-length-cong xs≈ys)) eq ⟩
+  _↔_.to (Fin-length-cong xs≈ys) (index q)  ≡⟨ sym $ index-commutes xs≈ys q ⟩∎
+  index (_↔_.to (xs≈ys z) q)                ∎
hunk ./Bag-equivalence.agda 593
-  abstract
-
-    -- If the equality type is proof irrelevant (so that p and q are
-    -- equal), then this lemma can be proved without the help of
-    -- index-equality-preserved.
-
-    lemma : ∀ {xs ys} (inv : x ∷ xs ≈-bag x ∷ ys) →
-            Well-behaved (_↔_.to (inv z))
-    lemma {xs} inv {b = z∈xs} {a = p} {a′ = q} hyp₁ hyp₂ = ⊎.inj₁≢inj₂ (
-      fzero                                ≡⟨ refl ⟩
-      index {xs = x ∷ xs} (inj₁ p)         ≡⟨ cong index $ sym $ to-from hyp₂ ⟩
-      index {xs = x ∷ xs} (from (inj₁ q))  ≡⟨ index-equality-preserved (inverse ∘ inv) refl ⟩
-      index {xs = x ∷ xs} (from (inj₁ p))  ≡⟨ cong index $ to-from hyp₁ ⟩
-      index {xs = x ∷ xs} (inj₂ z∈xs)      ≡⟨ refl ⟩∎
-      fsuc (index {xs = xs} z∈xs)          ∎)
-      where open _↔_ (inv z)
+  -- If the equality type is proof irrelevant (so that p and q are
+  -- equal), then this lemma can be proved without the help of
+  -- index-equality-preserved.
+
+  lemma : ∀ {xs ys} (inv : x ∷ xs ≈-bag x ∷ ys) →
+          Well-behaved (_↔_.to (inv z))
+  lemma {xs} inv {b = z∈xs} {a = p} {a′ = q} hyp₁ hyp₂ = ⊎.inj₁≢inj₂ (
+    fzero                                ≡⟨ refl ⟩
+    index {xs = x ∷ xs} (inj₁ p)         ≡⟨ cong index $ sym $ to-from hyp₂ ⟩
+    index {xs = x ∷ xs} (from (inj₁ q))  ≡⟨ index-equality-preserved (inverse ∘ inv) refl ⟩
+    index {xs = x ∷ xs} (from (inj₁ p))  ≡⟨ cong index $ to-from hyp₁ ⟩
+    index {xs = x ∷ xs} (inj₂ z∈xs)      ≡⟨ refl ⟩∎
+    fsuc (index {xs = xs} z∈xs)          ∎)
+    where open _↔_ (inv z)
hunk ./Bijection.agda 89
-  abstract
-    from∘to : ∀ x → from g (from f (to f (to g x))) ≡ x
-    from∘to = λ x →
-      from g (from f (to f (to g x)))  ≡⟨ cong (from g) (left-inverse-of f (to g x)) ⟩
-      from g (to g x)                  ≡⟨ left-inverse-of g x ⟩∎
-      x                                ∎
+  from∘to : ∀ x → from g (from f (to f (to g x))) ≡ x
+  from∘to = λ x →
+    from g (from f (to f (to g x)))  ≡⟨ cong (from g) (left-inverse-of f (to g x)) ⟩
+    from g (to g x)                  ≡⟨ left-inverse-of g x ⟩∎
+    x                                ∎
hunk ./Bijection.agda 145
-  abstract
-
-    to∘from : ∀ eq → to (from {p₁ = p₁} {p₂ = p₂} eq) ≡ eq
-    to∘from = elim (λ p≡q → to (from p≡q) ≡ p≡q) λ x →
-      let lem = subst-refl B (proj₂ x) in
-      to (from (refl x))                          ≡⟨ cong to (elim-refl from-P _) ⟩
-      to (refl (proj₁ x) , lem)                   ≡⟨ Σ-≡,≡→≡-reflˡ _ ⟩
-      cong (_,_ (proj₁ x)) (trans (sym lem) lem)  ≡⟨ cong (cong (_,_ (proj₁ x))) $ trans-symˡ lem ⟩
-      cong (_,_ (proj₁ x)) (refl (proj₂ x))       ≡⟨ cong-refl (_,_ (proj₁ x)) {x = proj₂ x} ⟩∎
-      refl x                                      ∎
-
-    from∘to : ∀ p → from (to {p₁ = p₁} {p₂ = p₂} p) ≡ p
-    from∘to p = elim
-      (λ {x₁ x₂} x₁≡x₂ →
-         ∀ {y₁ y₂} (y₁′≡y₂ : subst B x₁≡x₂ y₁ ≡ y₂) →
-         from (to (x₁≡x₂ , y₁′≡y₂)) ≡ (x₁≡x₂ , y₁′≡y₂))
-      (λ x {y₁} y₁′≡y₂ → elim
-         (λ {y₁ y₂} (y₁≡y₂ : y₁ ≡ y₂) →
-            (y₁′≡y₂ : subst B (refl x) y₁ ≡ y₂) →
-            y₁≡y₂ ≡ trans (sym $ subst-refl B y₁) y₁′≡y₂ →
-            from (to (refl x , y₁′≡y₂)) ≡ (refl x , y₁′≡y₂))
-         (λ y y′≡y eq →
-          let lem = subst-refl B y in
-          from (to (refl x , y′≡y))                   ≡⟨ cong from $ Σ-≡,≡→≡-reflˡ _ ⟩
-          from (cong (_,_ x) (trans (sym lem) y′≡y))  ≡⟨ cong (from ⊚ cong (_,_ x)) $ sym eq ⟩
-          from (cong (_,_ x) (refl y))                ≡⟨ cong from $ cong-refl (_,_ x) {x = y} ⟩
-          from (refl (x , y))                         ≡⟨ elim-refl from-P _ ⟩
-          (refl x , lem)                              ≡⟨ cong (_,_ (refl x)) (
-             lem                                           ≡⟨ sym $ trans-reflʳ _ ⟩
-             trans lem (refl _)                            ≡⟨ cong (trans lem) eq ⟩
-             trans lem (trans (sym lem) y′≡y)              ≡⟨ sym $ trans-assoc _ _ _ ⟩
-             trans (trans lem (sym lem)) y′≡y              ≡⟨ cong (λ p → trans p y′≡y) $ trans-symʳ lem ⟩
-             trans (refl _) y′≡y                           ≡⟨ trans-reflˡ _ ⟩∎
-             y′≡y                                          ∎) ⟩∎
-          (refl x , y′≡y)                             ∎)
-         (trans (sym $ subst-refl B y₁) y₁′≡y₂)
-         y₁′≡y₂
-         (refl _))
-      (proj₁ p) (proj₂ p)
+  to∘from : ∀ eq → to (from {p₁ = p₁} {p₂ = p₂} eq) ≡ eq
+  to∘from = elim (λ p≡q → to (from p≡q) ≡ p≡q) λ x →
+    let lem = subst-refl B (proj₂ x) in
+    to (from (refl x))                          ≡⟨ cong to (elim-refl from-P _) ⟩
+    to (refl (proj₁ x) , lem)                   ≡⟨ Σ-≡,≡→≡-reflˡ _ ⟩
+    cong (_,_ (proj₁ x)) (trans (sym lem) lem)  ≡⟨ cong (cong (_,_ (proj₁ x))) $ trans-symˡ lem ⟩
+    cong (_,_ (proj₁ x)) (refl (proj₂ x))       ≡⟨ cong-refl (_,_ (proj₁ x)) {x = proj₂ x} ⟩∎
+    refl x                                      ∎
+
+  from∘to : ∀ p → from (to {p₁ = p₁} {p₂ = p₂} p) ≡ p
+  from∘to p = elim
+    (λ {x₁ x₂} x₁≡x₂ →
+       ∀ {y₁ y₂} (y₁′≡y₂ : subst B x₁≡x₂ y₁ ≡ y₂) →
+       from (to (x₁≡x₂ , y₁′≡y₂)) ≡ (x₁≡x₂ , y₁′≡y₂))
+    (λ x {y₁} y₁′≡y₂ → elim
+       (λ {y₁ y₂} (y₁≡y₂ : y₁ ≡ y₂) →
+          (y₁′≡y₂ : subst B (refl x) y₁ ≡ y₂) →
+          y₁≡y₂ ≡ trans (sym $ subst-refl B y₁) y₁′≡y₂ →
+          from (to (refl x , y₁′≡y₂)) ≡ (refl x , y₁′≡y₂))
+       (λ y y′≡y eq →
+        let lem = subst-refl B y in
+        from (to (refl x , y′≡y))                   ≡⟨ cong from $ Σ-≡,≡→≡-reflˡ _ ⟩
+        from (cong (_,_ x) (trans (sym lem) y′≡y))  ≡⟨ cong (from ⊚ cong (_,_ x)) $ sym eq ⟩
+        from (cong (_,_ x) (refl y))                ≡⟨ cong from $ cong-refl (_,_ x) {x = y} ⟩
+        from (refl (x , y))                         ≡⟨ elim-refl from-P _ ⟩
+        (refl x , lem)                              ≡⟨ cong (_,_ (refl x)) (
+           lem                                           ≡⟨ sym $ trans-reflʳ _ ⟩
+           trans lem (refl _)                            ≡⟨ cong (trans lem) eq ⟩
+           trans lem (trans (sym lem) y′≡y)              ≡⟨ sym $ trans-assoc _ _ _ ⟩
+           trans (trans lem (sym lem)) y′≡y              ≡⟨ cong (λ p → trans p y′≡y) $ trans-symʳ lem ⟩
+           trans (refl _) y′≡y                           ≡⟨ trans-reflˡ _ ⟩∎
+           y′≡y                                          ∎) ⟩∎
+        (refl x , y′≡y)                             ∎)
+       (trans (sym $ subst-refl B y₁) y₁′≡y₂)
+       y₁′≡y₂
+       (refl _))
+    (proj₁ p) (proj₂ p)
hunk ./Bijection.agda 204
-  abstract
-
-    to∘from : ∀ ix≡iy → to (from ix≡iy) ≡ ix≡iy
-    to∘from ix≡iy =
-      cong inj₁ (⊎.cancel-inj₁ ix≡iy)              ≡⟨ cong-∘ inj₁ [ P.id , const x ] ix≡iy ⟩
-      cong f ix≡iy                                 ≡⟨ cong-roughly-id f p ix≡iy _ _ f≡id ⟩
-      trans (refl _) (trans ix≡iy $ sym (refl _))  ≡⟨ trans-reflˡ _ ⟩
-      trans ix≡iy (sym $ refl _)                   ≡⟨ cong (trans ix≡iy) sym-refl ⟩
-      trans ix≡iy (refl _)                         ≡⟨ trans-reflʳ _ ⟩∎
-      ix≡iy                                        ∎
-      where
-      f : A ⊎ B → A ⊎ B
-      f = inj₁ ⊚ [ P.id , const x ]
-
-      p : A ⊎ B → Bool
-      p = [ const true , const false ]
-
-      f≡id : ∀ z → T (p z) → f z ≡ z
-      f≡id (inj₁ x) = const (refl (inj₁ x))
-      f≡id (inj₂ y) = ⊥-elim
-
-    from∘to : ∀ x≡y → from (to x≡y) ≡ x≡y
-    from∘to x≡y =
-      cong [ P.id , const x ] (cong inj₁ x≡y)  ≡⟨ cong-∘ [ P.id , const x ] inj₁ _ ⟩
-      cong P.id x≡y                            ≡⟨ sym $ cong-id _ ⟩∎
-      x≡y                                      ∎
+  to∘from : ∀ ix≡iy → to (from ix≡iy) ≡ ix≡iy
+  to∘from ix≡iy =
+    cong inj₁ (⊎.cancel-inj₁ ix≡iy)              ≡⟨ cong-∘ inj₁ [ P.id , const x ] ix≡iy ⟩
+    cong f ix≡iy                                 ≡⟨ cong-roughly-id f p ix≡iy _ _ f≡id ⟩
+    trans (refl _) (trans ix≡iy $ sym (refl _))  ≡⟨ trans-reflˡ _ ⟩
+    trans ix≡iy (sym $ refl _)                   ≡⟨ cong (trans ix≡iy) sym-refl ⟩
+    trans ix≡iy (refl _)                         ≡⟨ trans-reflʳ _ ⟩∎
+    ix≡iy                                        ∎
+    where
+    f : A ⊎ B → A ⊎ B
+    f = inj₁ ⊚ [ P.id , const x ]
+
+    p : A ⊎ B → Bool
+    p = [ const true , const false ]
+
+    f≡id : ∀ z → T (p z) → f z ≡ z
+    f≡id (inj₁ x) = const (refl (inj₁ x))
+    f≡id (inj₂ y) = ⊥-elim
+
+  from∘to : ∀ x≡y → from (to x≡y) ≡ x≡y
+  from∘to x≡y =
+    cong [ P.id , const x ] (cong inj₁ x≡y)  ≡⟨ cong-∘ [ P.id , const x ] inj₁ _ ⟩
+    cong P.id x≡y                            ≡⟨ sym $ cong-id _ ⟩∎
+    x≡y                                      ∎
hunk ./Bijection.agda 247
-  abstract
-
-    to∘from : ∀ ix≡iy → to (from ix≡iy) ≡ ix≡iy
-    to∘from ix≡iy =
-      cong inj₂ (⊎.cancel-inj₂ ix≡iy)              ≡⟨ cong-∘ inj₂ [ const x , P.id ] ix≡iy ⟩
-      cong f ix≡iy                                 ≡⟨ cong-roughly-id f p ix≡iy _ _ f≡id ⟩
-      trans (refl _) (trans ix≡iy $ sym (refl _))  ≡⟨ trans-reflˡ _ ⟩
-      trans ix≡iy (sym $ refl _)                   ≡⟨ cong (trans ix≡iy) sym-refl ⟩
-      trans ix≡iy (refl _)                         ≡⟨ trans-reflʳ _ ⟩∎
-      ix≡iy                                        ∎
-      where
-      f : A ⊎ B → A ⊎ B
-      f = inj₂ ⊚ [ const x , P.id ]
-
-      p : A ⊎ B → Bool
-      p = [ const false , const true ]
-
-      f≡id : ∀ z → T (p z) → f z ≡ z
-      f≡id (inj₁ x) = ⊥-elim
-      f≡id (inj₂ y) = const (refl (inj₂ y))
-
-    from∘to : ∀ x≡y → from (to x≡y) ≡ x≡y
-    from∘to x≡y =
-      cong [ const x , P.id ] (cong inj₂ x≡y)  ≡⟨ cong-∘ [ const x , P.id ] inj₂ _ ⟩
-      cong P.id x≡y                            ≡⟨ sym $ cong-id _ ⟩∎
-      x≡y                                      ∎
+  to∘from : ∀ ix≡iy → to (from ix≡iy) ≡ ix≡iy
+  to∘from ix≡iy =
+    cong inj₂ (⊎.cancel-inj₂ ix≡iy)              ≡⟨ cong-∘ inj₂ [ const x , P.id ] ix≡iy ⟩
+    cong f ix≡iy                                 ≡⟨ cong-roughly-id f p ix≡iy _ _ f≡id ⟩
+    trans (refl _) (trans ix≡iy $ sym (refl _))  ≡⟨ trans-reflˡ _ ⟩
+    trans ix≡iy (sym $ refl _)                   ≡⟨ cong (trans ix≡iy) sym-refl ⟩
+    trans ix≡iy (refl _)                         ≡⟨ trans-reflʳ _ ⟩∎
+    ix≡iy                                        ∎
+    where
+    f : A ⊎ B → A ⊎ B
+    f = inj₂ ⊚ [ const x , P.id ]
+
+    p : A ⊎ B → Bool
+    p = [ const false , const true ]
+
+    f≡id : ∀ z → T (p z) → f z ≡ z
+    f≡id (inj₁ x) = ⊥-elim
+    f≡id (inj₂ y) = const (refl (inj₂ y))
+
+  from∘to : ∀ x≡y → from (to x≡y) ≡ x≡y
+  from∘to x≡y =
+    cong [ const x , P.id ] (cong inj₂ x≡y)  ≡⟨ cong-∘ [ const x , P.id ] inj₂ _ ⟩
+    cong P.id x≡y                            ≡⟨ sym $ cong-id _ ⟩∎
+    x≡y                                      ∎
hunk ./Category.agda 134
-  abstract
-
-    -- "Is-isomorphism f" is a proposition.
-
-    Is-isomorphism-propositional :
-      ∀ {X Y} (f : Hom X Y) →
-      Is-proposition (Is-isomorphism f)
-    Is-isomorphism-propositional f =
-      _⇔_.from propositional⇔irrelevant
-        λ { (g , fg , gf) (g′ , fg′ , g′f) →
-            Σ-≡,≡→≡ (g             ≡⟨ sym left-identity ⟩
-                     id ∙ g        ≡⟨ cong (λ h → h ∙ g) $ sym g′f ⟩
-                     (g′ ∙ f) ∙ g  ≡⟨ sym assoc ⟩
-                     g′ ∙ (f ∙ g)  ≡⟨ cong (_∙_ g′) fg ⟩
-                     g′ ∙ id       ≡⟨ right-identity ⟩∎
-                     g′            ∎)
-                    (Σ-≡,≡→≡ (_⇔_.to set⇔UIP Hom-is-set _ _)
-                             (_⇔_.to set⇔UIP Hom-is-set _ _)) }
-
-    -- Isomorphism equality is equivalent to "forward morphism"
-    -- equality.
-
-    ≡≃≡¹ : ∀ {X Y} {f g : X ≅ Y} → (f ≡ g) ≃ (f ¹ ≡ g ¹)
-    ≡≃≡¹ {f = f} {g} =
-      (f ≡ g)      ↔⟨ inverse $ ignore-propositional-component $ Is-isomorphism-propositional _ ⟩□
-      (f ¹ ≡ g ¹)  □
-
-    -- The type of isomorphisms (between two objects) is a set.
-
-    ≅-set : ∀ {X Y} → Is-set (X ≅ Y)
-    ≅-set = Σ-closure 2 Hom-is-set
-                      (λ _ → mono₁ 1 $ Is-isomorphism-propositional _)
+  -- "Is-isomorphism f" is a proposition.
+
+  Is-isomorphism-propositional :
+    ∀ {X Y} (f : Hom X Y) →
+    Is-proposition (Is-isomorphism f)
+  Is-isomorphism-propositional f =
+    _⇔_.from propositional⇔irrelevant
+      λ { (g , fg , gf) (g′ , fg′ , g′f) →
+          Σ-≡,≡→≡ (g             ≡⟨ sym left-identity ⟩
+                   id ∙ g        ≡⟨ cong (λ h → h ∙ g) $ sym g′f ⟩
+                   (g′ ∙ f) ∙ g  ≡⟨ sym assoc ⟩
+                   g′ ∙ (f ∙ g)  ≡⟨ cong (_∙_ g′) fg ⟩
+                   g′ ∙ id       ≡⟨ right-identity ⟩∎
+                   g′            ∎)
+                  (Σ-≡,≡→≡ (_⇔_.to set⇔UIP Hom-is-set _ _)
+                           (_⇔_.to set⇔UIP Hom-is-set _ _)) }
+
+  -- Isomorphism equality is equivalent to "forward morphism"
+  -- equality.
+
+  ≡≃≡¹ : ∀ {X Y} {f g : X ≅ Y} → (f ≡ g) ≃ (f ¹ ≡ g ¹)
+  ≡≃≡¹ {f = f} {g} =
+    (f ≡ g)      ↔⟨ inverse $ ignore-propositional-component $ Is-isomorphism-propositional _ ⟩□
+    (f ¹ ≡ g ¹)  □
+
+  -- The type of isomorphisms (between two objects) is a set.
+
+  ≅-set : ∀ {X Y} → Is-set (X ≅ Y)
+  ≅-set = Σ-closure 2 Hom-is-set
+                    (λ _ → mono₁ 1 $ Is-isomorphism-propositional _)
hunk ./Category.agda 177
-    abstract
-      fg : ∀ {X Y Z} (f : X ≅ Y) (g : Y ≅ Z) →
-           (f ¹ ∙ g ¹) ∙ (g ⁻¹ ∙ f ⁻¹) ≡ id
-      fg f g =
-        (f ¹ ∙ g ¹) ∙ (g ⁻¹ ∙ f ⁻¹)  ≡⟨ sym assoc ⟩
-        f ¹ ∙ (g ¹ ∙ (g ⁻¹ ∙ f ⁻¹))  ≡⟨ cong (_∙_ (f ¹)) assoc ⟩
-        f ¹ ∙ ((g ¹ ∙ g ⁻¹) ∙ f ⁻¹)  ≡⟨ cong (λ h → f ¹ ∙ (h ∙ f ⁻¹)) $ g ¹⁻¹ ⟩
-        f ¹ ∙ (id ∙ f ⁻¹)            ≡⟨ cong (_∙_ (f ¹)) left-identity ⟩
-        f ¹ ∙ f ⁻¹                   ≡⟨ f ¹⁻¹ ⟩∎
-        id                           ∎
-
-      gf : ∀ {X Y Z} (f : X ≅ Y) (g : Y ≅ Z) →
-           (g ⁻¹ ∙ f ⁻¹) ∙ (f ¹ ∙ g ¹) ≡ id
-      gf f g =
-        (g ⁻¹ ∙ f ⁻¹) ∙ (f ¹ ∙ g ¹)  ≡⟨ sym assoc ⟩
-        g ⁻¹ ∙ (f ⁻¹ ∙ (f ¹ ∙ g ¹))  ≡⟨ cong (_∙_ (g ⁻¹)) assoc ⟩
-        g ⁻¹ ∙ ((f ⁻¹ ∙ f ¹) ∙ g ¹)  ≡⟨ cong (λ h → g ⁻¹ ∙ (h ∙ g ¹)) $ f ⁻¹¹ ⟩
-        g ⁻¹ ∙ (id ∙ g ¹)            ≡⟨ cong (_∙_ (g ⁻¹)) left-identity ⟩
-        g ⁻¹ ∙ g ¹                   ≡⟨ g ⁻¹¹ ⟩∎
-        id                           ∎
+    fg : ∀ {X Y Z} (f : X ≅ Y) (g : Y ≅ Z) →
+         (f ¹ ∙ g ¹) ∙ (g ⁻¹ ∙ f ⁻¹) ≡ id
+    fg f g =
+      (f ¹ ∙ g ¹) ∙ (g ⁻¹ ∙ f ⁻¹)  ≡⟨ sym assoc ⟩
+      f ¹ ∙ (g ¹ ∙ (g ⁻¹ ∙ f ⁻¹))  ≡⟨ cong (_∙_ (f ¹)) assoc ⟩
+      f ¹ ∙ ((g ¹ ∙ g ⁻¹) ∙ f ⁻¹)  ≡⟨ cong (λ h → f ¹ ∙ (h ∙ f ⁻¹)) $ g ¹⁻¹ ⟩
+      f ¹ ∙ (id ∙ f ⁻¹)            ≡⟨ cong (_∙_ (f ¹)) left-identity ⟩
+      f ¹ ∙ f ⁻¹                   ≡⟨ f ¹⁻¹ ⟩∎
+      id                           ∎
+
+    gf : ∀ {X Y Z} (f : X ≅ Y) (g : Y ≅ Z) →
+         (g ⁻¹ ∙ f ⁻¹) ∙ (f ¹ ∙ g ¹) ≡ id
+    gf f g =
+      (g ⁻¹ ∙ f ⁻¹) ∙ (f ¹ ∙ g ¹)  ≡⟨ sym assoc ⟩
+      g ⁻¹ ∙ (f ⁻¹ ∙ (f ¹ ∙ g ¹))  ≡⟨ cong (_∙_ (g ⁻¹)) assoc ⟩
+      g ⁻¹ ∙ ((f ⁻¹ ∙ f ¹) ∙ g ¹)  ≡⟨ cong (λ h → g ⁻¹ ∙ (h ∙ g ¹)) $ f ⁻¹¹ ⟩
+      g ⁻¹ ∙ (id ∙ g ¹)            ≡⟨ cong (_∙_ (g ⁻¹)) left-identity ⟩
+      g ⁻¹ ∙ g ¹                   ≡⟨ g ⁻¹¹ ⟩∎
+      id                           ∎
hunk ./Category.agda 361
-    abstract
-
-      is-equiv : ∀ {X Y} → Is-equivalence (P≅.≡→≅ {X = X} {Y = Y})
-      is-equiv (X≅Y , X≅Y-iso) =
-        Σ-map (Σ-map
-          P.id
-          (λ {X≡Y} ≡→≅[X≡Y]≡X≅Y →
-             elim (λ {X Y} X≡Y →
-                     (X≅Y : X ≅ Y) (X≅Y-iso : P≅.Is-isomorphism X≅Y) →
-                     ≡→≅ X≡Y ≡ X≅Y →
-                     P≅.≡→≅ X≡Y ≡ (X≅Y , X≅Y-iso))
-                  (λ X X≅X X≅X-iso ≡→≅[refl]≡X≅X →
-                     P≅.≡→≅ (refl X)  ≡⟨ P≅.≡→≅-refl ⟩
-                     P≅.id≅           ≡⟨ Σ-≡,≡→≡ (id≅           ≡⟨ sym ≡→≅-refl ⟩
-                                                  ≡→≅ (refl X)  ≡⟨ ≡→≅[refl]≡X≅X ⟩∎
-                                                  X≅X           ∎)
-                                                 (_⇔_.to propositional⇔irrelevant (P≅.Is-isomorphism-propositional _) _ _) ⟩∎
-                     (X≅X , X≅X-iso)  ∎)
-                  X≡Y X≅Y X≅Y-iso
-                  ≡→≅[X≡Y]≡X≅Y))
-          (λ { {X≡Y , _} ∀y→≡y → λ { (X≡Y′ , ≡→≅[X≡Y′]≡X≅Y) →
-             let lemma =
-                   ≡→≅ X≡Y′             ≡⟨ elim (λ X≡Y′ → ≡→≅ X≡Y′ ≡ proj₁ (P≅.≡→≅ X≡Y′))
-                                                (λ X → ≡→≅ (refl X)             ≡⟨ ≡→≅-refl ⟩
-                                                       id≅                      ≡⟨ cong proj₁ $ sym P≅.≡→≅-refl ⟩∎
-                                                       proj₁ (P≅.≡→≅ (refl X))  ∎)
-                                                X≡Y′ ⟩
-                   proj₁ (P≅.≡→≅ X≡Y′)  ≡⟨ cong proj₁ ≡→≅[X≡Y′]≡X≅Y ⟩∎
-                   X≅Y                  ∎ in
-             (X≡Y , _)   ≡⟨ Σ-≡,≡→≡ (cong proj₁ (∀y→≡y (X≡Y′ , lemma)))
-                                    (_⇔_.to set⇔UIP P≅.≅-set _ _) ⟩∎
-             (X≡Y′ , _)  ∎ } })
-          (≡→≅-equivalence X≅Y)
+    is-equiv : ∀ {X Y} → Is-equivalence (P≅.≡→≅ {X = X} {Y = Y})
+    is-equiv (X≅Y , X≅Y-iso) =
+      Σ-map (Σ-map
+        P.id
+        (λ {X≡Y} ≡→≅[X≡Y]≡X≅Y →
+           elim (λ {X Y} X≡Y →
+                   (X≅Y : X ≅ Y) (X≅Y-iso : P≅.Is-isomorphism X≅Y) →
+                   ≡→≅ X≡Y ≡ X≅Y →
+                   P≅.≡→≅ X≡Y ≡ (X≅Y , X≅Y-iso))
+                (λ X X≅X X≅X-iso ≡→≅[refl]≡X≅X →
+                   P≅.≡→≅ (refl X)  ≡⟨ P≅.≡→≅-refl ⟩
+                   P≅.id≅           ≡⟨ Σ-≡,≡→≡ (id≅           ≡⟨ sym ≡→≅-refl ⟩
+                                                ≡→≅ (refl X)  ≡⟨ ≡→≅[refl]≡X≅X ⟩∎
+                                                X≅X           ∎)
+                                               (_⇔_.to propositional⇔irrelevant (P≅.Is-isomorphism-propositional _) _ _) ⟩∎
+                   (X≅X , X≅X-iso)  ∎)
+                X≡Y X≅Y X≅Y-iso
+                ≡→≅[X≡Y]≡X≅Y))
+        (λ { {X≡Y , _} ∀y→≡y → λ { (X≡Y′ , ≡→≅[X≡Y′]≡X≅Y) →
+           let lemma =
+                 ≡→≅ X≡Y′             ≡⟨ elim (λ X≡Y′ → ≡→≅ X≡Y′ ≡ proj₁ (P≅.≡→≅ X≡Y′))
+                                              (λ X → ≡→≅ (refl X)             ≡⟨ ≡→≅-refl ⟩
+                                                     id≅                      ≡⟨ cong proj₁ $ sym P≅.≡→≅-refl ⟩∎
+                                                     proj₁ (P≅.≡→≅ (refl X))  ∎)
+                                              X≡Y′ ⟩
+                 proj₁ (P≅.≡→≅ X≡Y′)  ≡⟨ cong proj₁ ≡→≅[X≡Y′]≡X≅Y ⟩∎
+                 X≅Y                  ∎ in
+           (X≡Y , _)   ≡⟨ Σ-≡,≡→≡ (cong proj₁ (∀y→≡y (X≡Y′ , lemma)))
+                                  (_⇔_.to set⇔UIP P≅.≅-set _ _) ⟩∎
+           (X≡Y′ , _)  ∎ } })
+        (≡→≅-equivalence X≅Y)
hunk ./Category.agda 454
-  abstract
-
-    -- _≡_ and _≅_ are pointwise equivalent…
-
-    cong-Type : {X Y : Obj} → (X ≡ Y) ≃ (Type X ≡ Type Y)
-    cong-Type = Eq.↔⇒≃ $ inverse $
-      ignore-propositional-component (H-level-propositional ext 2)
-
-    ≃≃≅ : (X Y : Obj) → (Type X ≃ Type Y) ≃ (X ≅ Y)
-    ≃≃≅ = ≃≃≅-Set ℓ ext
-
-    ≡≃≅ : ∀ {X Y} → (X ≡ Y) ≃ (X ≅ Y)
-    ≡≃≅ {X} {Y} = ≃≃≅ X Y ⊚ ≡≃≃ univ ⊚ cong-Type
-
-    -- …and the proof maps reflexivity to the identity isomorphism.
-
-    ≡≃≅-refl-is-id≅ : ∀ {X} → _≃_.to ≡≃≅ (refl X) ≡ id≅ {X = X}
-    ≡≃≅-refl-is-id≅ {X} =
-      _≃_.to (≃≃≅ X X) (≡⇒≃ (proj₁ (Σ-≡,≡←≡ (refl X))))  ≡⟨ cong (_≃_.to (≃≃≅ X X) ∘ ≡⇒≃ ∘ proj₁) Σ-≡,≡←≡-refl ⟩
-      _≃_.to (≃≃≅ X X) (≡⇒≃ (refl (Type X)))             ≡⟨ cong (_≃_.to (≃≃≅ X X)) ≡⇒≃-refl ⟩
-      _≃_.to (≃≃≅ X X) Eq.id                             ≡⟨ _≃_.from (≡≃≡¹ {X = X} {Y = X}) (refl P.id) ⟩∎
-      id≅ {X = X}                                        ∎
+  -- _≡_ and _≅_ are pointwise equivalent…
+
+  cong-Type : {X Y : Obj} → (X ≡ Y) ≃ (Type X ≡ Type Y)
+  cong-Type = Eq.↔⇒≃ $ inverse $
+    ignore-propositional-component (H-level-propositional ext 2)
+
+  ≃≃≅ : (X Y : Obj) → (Type X ≃ Type Y) ≃ (X ≅ Y)
+  ≃≃≅ = ≃≃≅-Set ℓ ext
+
+  ≡≃≅ : ∀ {X Y} → (X ≡ Y) ≃ (X ≅ Y)
+  ≡≃≅ {X} {Y} = ≃≃≅ X Y ⊚ ≡≃≃ univ ⊚ cong-Type
+
+  -- …and the proof maps reflexivity to the identity isomorphism.
+
+  ≡≃≅-refl-is-id≅ : ∀ {X} → _≃_.to ≡≃≅ (refl X) ≡ id≅ {X = X}
+  ≡≃≅-refl-is-id≅ {X} =
+    _≃_.to (≃≃≅ X X) (≡⇒≃ (proj₁ (Σ-≡,≡←≡ (refl X))))  ≡⟨ cong (_≃_.to (≃≃≅ X X) ∘ ≡⇒≃ ∘ proj₁) Σ-≡,≡←≡-refl ⟩
+    _≃_.to (≃≃≅ X X) (≡⇒≃ (refl (Type X)))             ≡⟨ cong (_≃_.to (≃≃≅ X X)) ≡⇒≃-refl ⟩
+    _≃_.to (≃≃≅ X X) Eq.id                             ≡⟨ _≃_.from (≡≃≡¹ {X = X} {Y = X}) (refl P.id) ⟩∎
+    id≅ {X = X}                                        ∎
hunk ./Container.agda 463
-    abstract
-      irr : f-eq′ ≡ f-eq
-      irr = proj₁ $
-        Eq.propositional (lower-extensionality a a ext) f _ _
-
-      f≡f : ⟨ f , f-eq′ ⟩ ≡ ⟨ f , f-eq ⟩
-      f≡f = cong (⟨_,_⟩ f) irr
-
-      cong-to-f≡f : cong _≃_.to f≡f ≡ refl {x = f}
-      cong-to-f≡f =
-        cong _≃_.to f≡f              ≡⟨ cong-∘ _≃_.to (⟨_,_⟩ f) irr ⟩
-        cong (_≃_.to ∘ ⟨_,_⟩ f) irr  ≡⟨ cong-const irr ⟩∎
-        refl                         ∎
+    irr : f-eq′ ≡ f-eq
+    irr = proj₁ $
+      Eq.propositional (lower-extensionality a a ext) f _ _
+
+    f≡f : ⟨ f , f-eq′ ⟩ ≡ ⟨ f , f-eq ⟩
+    f≡f = cong (⟨_,_⟩ f) irr
+
+    cong-to-f≡f : cong _≃_.to f≡f ≡ refl {x = f}
+    cong-to-f≡f =
+      cong _≃_.to f≡f              ≡⟨ cong-∘ _≃_.to (⟨_,_⟩ f) irr ⟩
+      cong (_≃_.to ∘ ⟨_,_⟩ f) irr  ≡⟨ cong-const irr ⟩∎
+      refl                         ∎
hunk ./Equality.agda 121
-  abstract
-
-    -- "Evaluation rule" for cong.
-
-    cong-refl : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) {x : A} →
-                cong f (refl x) ≡ refl (f x)
-    cong-refl f = elim-refl (λ {u v} _ → f u ≡ f v) (refl ∘ f)
+  -- "Evaluation rule" for cong.
+
+  cong-refl : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) {x : A} →
+              cong f (refl x) ≡ refl (f x)
+  cong-refl f = elim-refl (λ {u v} _ → f u ≡ f v) (refl ∘ f)
hunk ./Equality.agda 157
-  abstract
-
-    -- "Evaluation rule" for singleton-contractible.
-
-    singleton-contractible-refl :
-      ∀ {a} {A : Set a} (x : A) →
-      proj₂ (singleton-contractible x) (x , refl x) ≡ refl (x , refl x)
-    singleton-contractible-refl x =
-      elim-refl (λ {u v} u≡v → _≡_ {A = Singleton v}
-                                   (v , refl v) (u , u≡v))
-                _
+  -- "Evaluation rule" for singleton-contractible.
+
+  singleton-contractible-refl :
+    ∀ {a} {A : Set a} (x : A) →
+    proj₂ (singleton-contractible x) (x , refl x) ≡ refl (x , refl x)
+  singleton-contractible-refl x =
+    elim-refl (λ {u v} u≡v → _≡_ {A = Singleton v}
+                                 (v , refl v) (u , u≡v))
+              _
hunk ./Equality.agda 210
-  abstract
-
-    -- Congruence.
-
-    cong : ∀ {a b} {A : Set a} {B : Set b}
-           (f : A → B) {x y : A} → x ≡ y → f x ≡ f y
-    cong f {x} x≡y =
-      subst (λ y → x ≡ y → f x ≡ f y) x≡y (λ _ → refl (f x)) x≡y
+  -- Congruence.
+
+  cong : ∀ {a b} {A : Set a} {B : Set b}
+         (f : A → B) {x y : A} → x ≡ y → f x ≡ f y
+  cong f {x} x≡y =
+    subst (λ y → x ≡ y → f x ≡ f y) x≡y (λ _ → refl (f x)) x≡y
hunk ./Equality.agda 222
-  abstract
-
-    -- "Evaluation rule" for sym.
-
-    sym-refl : ∀ {a} {A : Set a} {x : A} → sym (refl x) ≡ refl x
-    sym-refl {x = x} =
-      cong (λ f → f (refl x)) $
-        subst-refl (λ z → x ≡ z → z ≡ x) id
+  -- "Evaluation rule" for sym.
+
+  sym-refl : ∀ {a} {A : Set a} {x : A} → sym (refl x) ≡ refl x
+  sym-refl {x = x} =
+    cong (λ f → f (refl x)) $
+      subst-refl (λ z → x ≡ z → z ≡ x) id
hunk ./Equality.agda 234
-  abstract
-
-    -- "Evaluation rule" for trans.
-
-    trans-refl-refl : ∀ {a} {A : Set a} {x : A} →
-                      trans (refl x) (refl x) ≡ refl x
-    trans-refl-refl {x = x} = subst-refl (_≡_ x) (refl x)
+  -- "Evaluation rule" for trans.
+
+  trans-refl-refl : ∀ {a} {A : Set a} {x : A} →
+                    trans (refl x) (refl x) ≡ refl x
+  trans-refl-refl {x = x} = subst-refl (_≡_ x) (refl x)
hunk ./Equality.agda 256
-  abstract
-
-    -- The J rule.
-
-    elim : ∀ {a p} {A : Set a} (P : {x y : A} → x ≡ y → Set p) →
-           (∀ x → P (refl x)) →
-           ∀ {x y} (x≡y : x ≡ y) → P x≡y
-    elim P p {x} {y} x≡y =
-      let lemma = proj₂ (singleton-contractible y) in
-      subst {A = Singleton y}
-            (P ∘ proj₂)
-            ((y , refl y)                      ≡⟨ sym (lemma (y , refl y)) ⟩
-             proj₁ (singleton-contractible y)  ≡⟨ lemma (x , x≡y) ⟩∎
-             (x , x≡y)                         ∎)
-            (p y)
-
-    -- Transitivity and symmetry sometimes cancel each other out.
-
-    trans-sym : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
-                trans (sym x≡y) x≡y ≡ refl y
-    trans-sym =
-      elim (λ {x y} (x≡y : x ≡ y) → trans (sym x≡y) x≡y ≡ refl y)
-           (λ x → trans (sym (refl x)) (refl x)  ≡⟨ cong (λ p → trans p (refl x)) sym-refl ⟩
-                  trans (refl x) (refl x)        ≡⟨ trans-refl-refl ⟩∎
-                  refl x                         ∎)
-
-    -- "Evaluation rule" for elim.
-
-    elim-refl : ∀ {a p} {A : Set a} (P : {x y : A} → x ≡ y → Set p)
-                (p : ∀ x → P (refl x)) {x} →
-                elim P p (refl x) ≡ p x
-    elim-refl P p {x} =
-      let lemma = proj₂ (singleton-contractible x) (x , refl x) in
-      subst {A = Singleton x} (P ∘ proj₂) (trans (sym lemma) lemma) (p x)  ≡⟨ cong (λ q → subst {A = Singleton x} (P ∘ proj₂) q (p x))
-                                                                                   (trans-sym lemma) ⟩
-      subst {A = Singleton x} (P ∘ proj₂) (refl (x , refl x))       (p x)  ≡⟨ subst-refl {A = Singleton x} (P ∘ proj₂) (p x) ⟩∎
-      p x                                                                  ∎
+  -- The J rule.
+
+  elim : ∀ {a p} {A : Set a} (P : {x y : A} → x ≡ y → Set p) →
+         (∀ x → P (refl x)) →
+         ∀ {x y} (x≡y : x ≡ y) → P x≡y
+  elim P p {x} {y} x≡y =
+    let lemma = proj₂ (singleton-contractible y) in
+    subst {A = Singleton y}
+          (P ∘ proj₂)
+          ((y , refl y)                      ≡⟨ sym (lemma (y , refl y)) ⟩
+           proj₁ (singleton-contractible y)  ≡⟨ lemma (x , x≡y) ⟩∎
+           (x , x≡y)                         ∎)
+          (p y)
+
+  -- Transitivity and symmetry sometimes cancel each other out.
+
+  trans-sym : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
+              trans (sym x≡y) x≡y ≡ refl y
+  trans-sym =
+    elim (λ {x y} (x≡y : x ≡ y) → trans (sym x≡y) x≡y ≡ refl y)
+         (λ x → trans (sym (refl x)) (refl x)  ≡⟨ cong (λ p → trans p (refl x)) sym-refl ⟩
+                trans (refl x) (refl x)        ≡⟨ trans-refl-refl ⟩∎
+                refl x                         ∎)
+
+  -- "Evaluation rule" for elim.
+
+  elim-refl : ∀ {a p} {A : Set a} (P : {x y : A} → x ≡ y → Set p)
+              (p : ∀ x → P (refl x)) {x} →
+              elim P p (refl x) ≡ p x
+  elim-refl P p {x} =
+    let lemma = proj₂ (singleton-contractible x) (x , refl x) in
+    subst {A = Singleton x} (P ∘ proj₂) (trans (sym lemma) lemma) (p x)  ≡⟨ cong (λ q → subst {A = Singleton x} (P ∘ proj₂) q (p x))
+                                                                                 (trans-sym lemma) ⟩
+    subst {A = Singleton x} (P ∘ proj₂) (refl (x , refl x))       (p x)  ≡⟨ subst-refl {A = Singleton x} (P ∘ proj₂) (p x) ⟩∎
+    p x                                                                  ∎
hunk ./Equality.agda 354
-  abstract
-
-    -- "Evaluation rule" for elim₁.
-
-    elim₁-refl : ∀ {a p} {A : Set a} {y : A}
-                 (P : ∀ {x} → x ≡ y → Set p) (p : P (refl y)) →
-                 elim₁ P p (refl y) ≡ p
-    elim₁-refl {y = y} P p =
-      subst {A = Singleton y} (P ∘ proj₂)
-            (proj₂ (singleton-contractible y) (y , refl y)) p    ≡⟨ cong (λ q → subst {A = Singleton y} (P ∘ proj₂) q p)
-                                                                         (singleton-contractible-refl y) ⟩
-      subst {A = Singleton y} (P ∘ proj₂) (refl (y , refl y)) p  ≡⟨ subst-refl {A = Singleton y} (P ∘ proj₂) p ⟩∎
-      p                                                          ∎
+  -- "Evaluation rule" for elim₁.
+
+  elim₁-refl : ∀ {a p} {A : Set a} {y : A}
+               (P : ∀ {x} → x ≡ y → Set p) (p : P (refl y)) →
+               elim₁ P p (refl y) ≡ p
+  elim₁-refl {y = y} P p =
+    subst {A = Singleton y} (P ∘ proj₂)
+          (proj₂ (singleton-contractible y) (y , refl y)) p    ≡⟨ cong (λ q → subst {A = Singleton y} (P ∘ proj₂) q p)
+                                                                       (singleton-contractible-refl y) ⟩
+    subst {A = Singleton y} (P ∘ proj₂) (refl (y , refl y)) p  ≡⟨ subst-refl {A = Singleton y} (P ∘ proj₂) p ⟩∎
+    p                                                          ∎
hunk ./Equality.agda 385
-  abstract
-
-    -- "Evaluation rule" for other-singleton-contractible.
-
-    other-singleton-contractible-refl :
-      ∀ {a} {A : Set a} (x : A) →
-      proj₂ (other-singleton-contractible x) (x , refl x) ≡
-      refl (x , refl x)
-    other-singleton-contractible-refl x =
-      elim-refl (λ {u v} u≡v → _≡_ {A = Other-singleton u}
-                                   (u , refl u) (v , u≡v))
-                _
+  -- "Evaluation rule" for other-singleton-contractible.
+
+  other-singleton-contractible-refl :
+    ∀ {a} {A : Set a} (x : A) →
+    proj₂ (other-singleton-contractible x) (x , refl x) ≡
+    refl (x , refl x)
+  other-singleton-contractible-refl x =
+    elim-refl (λ {u v} u≡v → _≡_ {A = Other-singleton u}
+                                 (u , refl u) (v , u≡v))
+              _
hunk ./Equality.agda 407
-  abstract
-
-    -- "Evaluation rule" for elim¹.
-
-    elim¹-refl : ∀ {a p} {A : Set a} {x : A}
-                 (P : ∀ {y} → x ≡ y → Set p) (p : P (refl x)) →
-                 elim¹ P p (refl x) ≡ p
-    elim¹-refl {x = x} P p =
-      subst {A = Other-singleton x} (P ∘ proj₂)
-            (proj₂ (other-singleton-contractible x) (x , refl x)) p    ≡⟨ cong (λ q → subst {A = Other-singleton x} (P ∘ proj₂) q p)
-                                                                               (other-singleton-contractible-refl x) ⟩
-      subst {A = Other-singleton x} (P ∘ proj₂) (refl (x , refl x)) p  ≡⟨ subst-refl {A = Other-singleton x} (P ∘ proj₂) p ⟩∎
-      p                                                                ∎
+  -- "Evaluation rule" for elim¹.
+
+  elim¹-refl : ∀ {a p} {A : Set a} {x : A}
+               (P : ∀ {y} → x ≡ y → Set p) (p : P (refl x)) →
+               elim¹ P p (refl x) ≡ p
+  elim¹-refl {x = x} P p =
+    subst {A = Other-singleton x} (P ∘ proj₂)
+          (proj₂ (other-singleton-contractible x) (x , refl x)) p    ≡⟨ cong (λ q → subst {A = Other-singleton x} (P ∘ proj₂) q p)
+                                                                             (other-singleton-contractible-refl x) ⟩
+    subst {A = Other-singleton x} (P ∘ proj₂) (refl (x , refl x)) p  ≡⟨ subst-refl {A = Other-singleton x} (P ∘ proj₂) p ⟩∎
+    p                                                                ∎
hunk ./Equality.agda 430
-  abstract
-
-    -- "Evaluation rule" for dependent-cong.
-
-    dependent-cong-refl :
-      ∀ {a b} {A : Set a} {B : A → Set b} (f : (x : A) → B x) {x} →
-      dependent-cong f (refl x) ≡ subst-refl B (f x)
-    dependent-cong-refl {B = B} f {x} =
-      elim (λ {x y} (x≡y : x ≡ y) → subst B x≡y (f x) ≡ f y)
-           (λ x → subst-refl B (f x))
-           (refl x)                                           ≡⟨ elim-refl _ _ ⟩∎
-      subst-refl B (f x)                                      ∎
+  -- "Evaluation rule" for dependent-cong.
+
+  dependent-cong-refl :
+    ∀ {a b} {A : Set a} {B : A → Set b} (f : (x : A) → B x) {x} →
+    dependent-cong f (refl x) ≡ subst-refl B (f x)
+  dependent-cong-refl {B = B} f {x} =
+    elim (λ {x y} (x≡y : x ≡ y) → subst B x≡y (f x) ≡ f y)
+         (λ x → subst-refl B (f x))
+         (refl x)                                           ≡⟨ elim-refl _ _ ⟩∎
+    subst-refl B (f x)                                      ∎
hunk ./Equality.agda 451
-  abstract
-
-    -- "Evaluation rule" for cong₂.
-
-    cong₂-refl : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
-                 (f : A → B → C) {x : A} {y : B} →
-                 cong₂ f (refl x) (refl y) ≡ refl (f x y)
-    cong₂-refl f {x} {y} =
-      trans (cong (flip f y) (refl x)) (cong (f x) (refl y))  ≡⟨ cong₂ trans (cong-refl (flip f y)) (cong-refl (f x)) ⟩
-      trans (refl (f x y)) (refl (f x y))                     ≡⟨ trans-refl-refl ⟩∎
-      refl (f x y)                                            ∎
+  -- "Evaluation rule" for cong₂.
+
+  cong₂-refl : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
+               (f : A → B → C) {x : A} {y : B} →
+               cong₂ f (refl x) (refl y) ≡ refl (f x y)
+  cong₂-refl f {x} {y} =
+    trans (cong (flip f y) (refl x)) (cong (f x) (refl y))  ≡⟨ cong₂ trans (cong-refl (flip f y)) (cong-refl (f x)) ⟩
+    trans (refl (f x y)) (refl (f x y))                     ≡⟨ trans-refl-refl ⟩∎
+    refl (f x y)                                            ∎
hunk ./Equality.agda 500
-  abstract
-
-    -- Extensionality at given levels works at lower levels as well.
-
-    lower-extensionality :
-      ∀ {a} â {b} b̂ →
-      Extensionality (a ⊔ â) (b ⊔ b̂) → Extensionality a b
-    lower-extensionality â b̂ ext f≡g =
-      cong (λ h → lower ∘ h ∘ lift) $
-        ext {A = ↑ â _} {B = ↑ b̂ ∘ _} (cong lift ∘ f≡g ∘ lower)
-
-    lower-extensionality₂ :
-      ∀ {a} {A : Set a} {b} b̂ →
-      ({B : A → Set (b ⊔ b̂)} → Extensionality′ A B) →
-      ({B : A → Set  b     } → Extensionality′ A B)
-    lower-extensionality₂ b̂ ext f≡g =
-      cong (λ h → lower ∘ h) $
-        ext {B = ↑ b̂ ∘ _} (cong lift ∘ f≡g)
-
-    -- Extensionality for explicit function types works for implicit
-    -- function types as well.
-
-    implicit-extensionality :
-      ∀ {a b} {A : Set a} {B : A → Set b} →
-      Extensionality′ A B →
-      {f g : {x : A} → B x} →
-      (∀ x → f {x} ≡ g {x}) → (λ {x} → f {x}) ≡ g
-    implicit-extensionality ext f≡g =
-      cong (λ f {x} → f x) $ ext f≡g
+  -- Extensionality at given levels works at lower levels as well.
+
+  lower-extensionality :
+    ∀ {a} â {b} b̂ →
+    Extensionality (a ⊔ â) (b ⊔ b̂) → Extensionality a b
+  lower-extensionality â b̂ ext f≡g =
+    cong (λ h → lower ∘ h ∘ lift) $
+      ext {A = ↑ â _} {B = ↑ b̂ ∘ _} (cong lift ∘ f≡g ∘ lower)
+
+  lower-extensionality₂ :
+    ∀ {a} {A : Set a} {b} b̂ →
+    ({B : A → Set (b ⊔ b̂)} → Extensionality′ A B) →
+    ({B : A → Set  b     } → Extensionality′ A B)
+  lower-extensionality₂ b̂ ext f≡g =
+    cong (λ h → lower ∘ h) $
+      ext {B = ↑ b̂ ∘ _} (cong lift ∘ f≡g)
+
+  -- Extensionality for explicit function types works for implicit
+  -- function types as well.
+
+  implicit-extensionality :
+    ∀ {a b} {A : Set a} {B : A → Set b} →
+    Extensionality′ A B →
+    {f g : {x : A} → B x} →
+    (∀ x → f {x} ≡ g {x}) → (λ {x} → f {x}) ≡ g
+  implicit-extensionality ext f≡g =
+    cong (λ f {x} → f x) $ ext f≡g
hunk ./Equality.agda 530
-  abstract
-
-    trans-reflʳ : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
-                  trans x≡y (refl y) ≡ x≡y
-    trans-reflʳ =
-      elim (λ {u v} u≡v → trans u≡v (refl v) ≡ u≡v)
-           (λ _ → trans-refl-refl)
-
-    trans-reflˡ : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
-                  trans (refl x) x≡y ≡ x≡y
-    trans-reflˡ =
-      elim (λ {u v} u≡v → trans (refl u) u≡v ≡ u≡v)
-           (λ _ → trans-refl-refl)
-
-    trans-assoc : ∀ {a} {A : Set a} {x y z u : A}
-                  (x≡y : x ≡ y) (y≡z : y ≡ z) (z≡u : z ≡ u) →
-                  trans (trans x≡y y≡z) z≡u ≡ trans x≡y (trans y≡z z≡u)
-    trans-assoc =
-      elim (λ x≡y → ∀ y≡z z≡u → trans (trans x≡y y≡z) z≡u ≡
-                                trans x≡y (trans y≡z z≡u))
-           (λ y y≡z z≡u →
-              trans (trans (refl y) y≡z) z≡u ≡⟨ cong₂ trans (trans-reflˡ y≡z) (refl z≡u) ⟩
-              trans y≡z z≡u                  ≡⟨ sym $ trans-reflˡ (trans y≡z z≡u) ⟩∎
-              trans (refl y) (trans y≡z z≡u) ∎)
-
-    sym-sym : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
-              sym (sym x≡y) ≡ x≡y
-    sym-sym = elim (λ {u v} u≡v → sym (sym u≡v) ≡ u≡v)
-                   (λ x → sym (sym (refl x))  ≡⟨ cong sym (sym-refl {x = x}) ⟩
-                          sym (refl x)        ≡⟨ sym-refl ⟩∎
-                          refl x              ∎)
-
-    sym-trans : ∀ {a} {A : Set a} {x y z : A}
-                (x≡y : x ≡ y) (y≡z : y ≡ z) →
-                sym (trans x≡y y≡z) ≡ trans (sym y≡z) (sym x≡y)
-    sym-trans {a} =
-      elim (λ x≡y → ∀ y≡z → sym (trans x≡y y≡z) ≡ trans (sym y≡z) (sym x≡y))
-           (λ y y≡z → sym (trans (refl y) y≡z)        ≡⟨ cong sym (trans-reflˡ y≡z) ⟩
-                      sym y≡z                         ≡⟨ sym $ trans-reflʳ (sym y≡z) ⟩
-                      trans (sym y≡z) (refl y)        ≡⟨ cong {a = a} {b = a} (trans (sym y≡z)) (sym sym-refl)  ⟩∎
-                      trans (sym y≡z) (sym (refl y))  ∎)
-
-    trans-symˡ : ∀ {a} {A : Set a} {x y : A} (p : x ≡ y) →
-                 trans (sym p) p ≡ refl y
-    trans-symˡ =
-      elim (λ p → trans (sym p) p ≡ refl _)
-           (λ x → trans (sym (refl x)) (refl x)  ≡⟨ trans-reflʳ _ ⟩
-                  sym (refl x)                   ≡⟨ sym-refl ⟩∎
-                  refl x                         ∎)
-
-    trans-symʳ : ∀ {a} {A : Set a} {x y : A} (p : x ≡ y) →
-                 trans p (sym p) ≡ refl _
-    trans-symʳ =
-      elim (λ p → trans p (sym p) ≡ refl _)
-           (λ x → trans (refl x) (sym (refl x))  ≡⟨ trans-reflˡ _ ⟩
-                  sym (refl x)                   ≡⟨ sym-refl ⟩∎
-                  refl x                         ∎)
-
-    cong-trans : ∀ {a b} {A : Set a} {B : Set b} {x y z : A}
-                 (f : A → B) (x≡y : x ≡ y) (y≡z : y ≡ z) →
-                 cong f (trans x≡y y≡z) ≡ trans (cong f x≡y) (cong f y≡z)
-    cong-trans f =
-      elim (λ x≡y → ∀ y≡z → cong f (trans x≡y y≡z) ≡
-                            trans (cong f x≡y) (cong f y≡z))
-           (λ y y≡z → cong f (trans (refl y) y≡z)           ≡⟨ cong (cong f) (trans-reflˡ _) ⟩
-                      cong f y≡z                            ≡⟨ sym $ trans-reflˡ (cong f y≡z) ⟩
-                      trans (refl (f y)) (cong f y≡z)       ≡⟨ cong₂ trans (sym (cong-refl f {x = y})) (refl (cong f y≡z)) ⟩∎
-                      trans (cong f (refl y)) (cong f y≡z)  ∎)
-
-    cong-id : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
-              x≡y ≡ cong id x≡y
-    cong-id = elim (λ u≡v → u≡v ≡ cong id u≡v)
-                   (λ x → refl x            ≡⟨ sym (cong-refl id {x = x}) ⟩∎
-                          cong id (refl x)  ∎)
-
-    cong-const : ∀ {a b} {A : Set a} {B : Set b} {x y : A} {z : B}
-                 (x≡y : x ≡ y) →
-                 cong (const z) x≡y ≡ refl z
-    cong-const {z = z} =
-      elim (λ u≡v → cong (const z) u≡v ≡ refl z)
-           (λ x → cong (const z) (refl x)  ≡⟨ cong-refl (const z) ⟩∎
-                  refl z                   ∎)
-
-    cong-∘ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} {x y : A}
-             (f : B → C) (g : A → B) (x≡y : x ≡ y) →
-             cong f (cong g x≡y) ≡ cong (f ∘ g) x≡y
-    cong-∘ f g = elim (λ x≡y → cong f (cong g x≡y) ≡ cong (f ∘ g) x≡y)
-                      (λ x → cong f (cong g (refl x))  ≡⟨ cong (cong f) (cong-refl g) ⟩
-                             cong f (refl (g x))       ≡⟨ cong-refl f ⟩
-                             refl (f (g x))            ≡⟨ sym (cong-refl (f ∘ g)) ⟩∎
-                             cong (f ∘ g) (refl x)     ∎)
-
-    cong-proj₁-cong₂-, :
-      ∀ {a b} {A : Set a} {B : Set b} {x y : A} {u v : B}
-      (x≡y : x ≡ y) (u≡v : u ≡ v) →
-      cong proj₁ (cong₂ _,_ x≡y u≡v) ≡ x≡y
-    cong-proj₁-cong₂-, {x = x} {y} {u} {v} x≡y u≡v =
-      cong proj₁ (trans (cong (flip _,_ u) x≡y) (cong (_,_ y) u≡v))  ≡⟨ cong-trans proj₁ _ _ ⟩
-
-      trans (cong proj₁ (cong (flip _,_ u) x≡y))
-            (cong proj₁ (cong (_,_ y) u≡v))                          ≡⟨ cong₂ trans (cong-∘ proj₁ (flip _,_ u) x≡y) (cong-∘ proj₁ (_,_ y) u≡v) ⟩
-
-      trans (cong id x≡y) (cong (const y) u≡v)                       ≡⟨ cong₂ trans (sym $ cong-id x≡y) (cong-const u≡v) ⟩
-
-      trans x≡y (refl y)                                             ≡⟨ trans-reflʳ x≡y ⟩∎
-
-      x≡y                                                            ∎
-
-    cong-proj₂-cong₂-, :
-      ∀ {a b} {A : Set a} {B : Set b} {x y : A} {u v : B}
-      (x≡y : x ≡ y) (u≡v : u ≡ v) →
-      cong proj₂ (cong₂ _,_ x≡y u≡v) ≡ u≡v
-    cong-proj₂-cong₂-, {x = x} {y} {u} {v} x≡y u≡v =
-      cong proj₂ (trans (cong (flip _,_ u) x≡y) (cong (_,_ y) u≡v))  ≡⟨ cong-trans proj₂ _ _ ⟩
-
-      trans (cong proj₂ (cong (flip _,_ u) x≡y))
-            (cong proj₂ (cong (_,_ y) u≡v))                          ≡⟨ cong₂ trans (cong-∘ proj₂ (flip _,_ u) x≡y) (cong-∘ proj₂ (_,_ y) u≡v) ⟩
-
-      trans (cong (const u) x≡y) (cong id u≡v)                       ≡⟨ cong₂ trans (cong-const x≡y) (sym $ cong-id u≡v) ⟩
-
-      trans (refl u) u≡v                                             ≡⟨ trans-reflˡ u≡v ⟩∎
-
-      u≡v                                                            ∎
-
-    cong-sym : ∀ {a b} {A : Set a} {B : Set b} {x y : A}
-               (f : A → B) (x≡y : x ≡ y) →
-               cong f (sym x≡y) ≡ sym (cong f x≡y)
-    cong-sym f = elim (λ x≡y → cong f (sym x≡y) ≡ sym (cong f x≡y))
-                      (λ x → cong f (sym (refl x))  ≡⟨ cong (cong f) sym-refl ⟩
-                             cong f (refl x)        ≡⟨ cong-refl f ⟩
-                             refl (f x)             ≡⟨ sym sym-refl ⟩
-                             sym (refl (f x))       ≡⟨ cong sym $ sym (cong-refl f {x = x}) ⟩∎
-                             sym (cong f (refl x))  ∎)
-
-    cong₂-reflˡ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
-                  (f : A → B → C) {x : A} {u v : B} {u≡v : u ≡ v} →
-                  cong₂ f (refl x) u≡v ≡ cong (f x) u≡v
-    cong₂-reflˡ f {x} {u} {u≡v = u≡v} =
-      trans (cong (flip f u) (refl x)) (cong (f x) u≡v)  ≡⟨ cong₂ trans (cong-refl (flip f u)) (refl _) ⟩
-      trans (refl (f x u)) (cong (f x) u≡v)              ≡⟨ trans-reflˡ _ ⟩∎
-      cong (f x) u≡v                                     ∎
-
-    cong₂-reflʳ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
-                  (f : A → B → C) {x y : A} {u : B} {x≡y : x ≡ y} →
-                  cong₂ f x≡y (refl u) ≡ cong (flip f u) x≡y
-    cong₂-reflʳ f {y = y} {u} {x≡y} =
-      trans (cong (flip f u) x≡y) (cong (f y) (refl u))  ≡⟨ cong (trans _) (cong-refl (f y)) ⟩
-      trans (cong (flip f u) x≡y) (refl (f y u))         ≡⟨ trans-reflʳ _ ⟩∎
-      cong (flip f u) x≡y                                ∎
-
-    cong₂-cong-cong :
-      ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d}
-        {x₁ x₂} {eq : x₁ ≡ x₂}
-      (f : A → B) (g : A → C) (h : B → C → D) →
-      cong₂ h (cong f eq) (cong g eq) ≡
-      cong (λ x → h (f x) (g x)) eq
-    cong₂-cong-cong f g h = elim¹
-      (λ eq → cong₂ h (cong f eq) (cong g eq) ≡
-              cong (λ x → h (f x) (g x)) eq)
-      (cong₂ h (cong f (refl _)) (cong g (refl _))  ≡⟨ cong₂ (cong₂ h) (cong-refl f) (cong-refl g) ⟩
-       cong₂ h (refl _) (refl _)                    ≡⟨ cong₂-refl h ⟩
-       refl _                                       ≡⟨ sym $ cong-refl (λ x → h (f x) (g x)) ⟩∎
-       cong (λ x → h (f x) (g x)) (refl _)          ∎)
-      _
-
-  abstract
-
-    elim-∘ :
-      ∀ {a p} {A : Set a} {x y : A}
-      (P Q : ∀ {x y} → x ≡ y → Set p)
-      (f : ∀ {x y} {x≡y : x ≡ y} → P x≡y → Q x≡y)
-      (r : ∀ x → P (refl x)) {x≡y : x ≡ y} →
-      f (elim P r x≡y) ≡ elim Q (f ∘ r) x≡y
-    elim-∘ {x = x} P Q f r {x≡y} = elim¹
-      (λ x≡y → f (elim P r x≡y) ≡
-               elim Q (f ∘ r) x≡y)
-      (f (elim P r (refl x))    ≡⟨ cong f $ elim-refl P _ ⟩
-       f (r x)                  ≡⟨ sym $ elim-refl Q _ ⟩∎
-       elim Q (f ∘ r) (refl x)  ∎)
-      x≡y
-
-    elim-cong :
-      ∀ {a b p} {A : Set a} {B : Set b} {x y : A}
-      (P : B → B → Set p) (f : A → B)
-      (r : ∀ x → P x x) {x≡y : x ≡ y} →
-      elim (λ {x y} _ → P x y) r (cong f x≡y) ≡
-      elim (λ {x y} _ → P (f x) (f y)) (r ∘ f) x≡y
-    elim-cong {x = x} P f r {x≡y} = elim¹
-      (λ x≡y → elim (λ {x y} _ → P x y) r (cong f x≡y) ≡
-               elim (λ {x y} _ → P (f x) (f y)) (r ∘ f) x≡y)
-      (elim (λ {x y} _ → P x y) r (cong f (refl x))       ≡⟨ cong (elim (λ {x y} _ → P x y) _) $ cong-refl f ⟩
-       elim (λ {x y} _ → P x y) r (refl (f x))            ≡⟨ elim-refl (λ {x y} _ → P x y) _ ⟩
-       r (f x)                                            ≡⟨ sym $ elim-refl (λ {x y} _ → P (f x) (f y)) _ ⟩∎
-       elim (λ {x y} _ → P (f x) (f y)) (r ∘ f) (refl x)  ∎)
-      x≡y
+  trans-reflʳ : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
+                trans x≡y (refl y) ≡ x≡y
+  trans-reflʳ =
+    elim (λ {u v} u≡v → trans u≡v (refl v) ≡ u≡v)
+         (λ _ → trans-refl-refl)
+
+  trans-reflˡ : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
+                trans (refl x) x≡y ≡ x≡y
+  trans-reflˡ =
+    elim (λ {u v} u≡v → trans (refl u) u≡v ≡ u≡v)
+         (λ _ → trans-refl-refl)
+
+  trans-assoc : ∀ {a} {A : Set a} {x y z u : A}
+                (x≡y : x ≡ y) (y≡z : y ≡ z) (z≡u : z ≡ u) →
+                trans (trans x≡y y≡z) z≡u ≡ trans x≡y (trans y≡z z≡u)
+  trans-assoc =
+    elim (λ x≡y → ∀ y≡z z≡u → trans (trans x≡y y≡z) z≡u ≡
+                              trans x≡y (trans y≡z z≡u))
+         (λ y y≡z z≡u →
+            trans (trans (refl y) y≡z) z≡u ≡⟨ cong₂ trans (trans-reflˡ y≡z) (refl z≡u) ⟩
+            trans y≡z z≡u                  ≡⟨ sym $ trans-reflˡ (trans y≡z z≡u) ⟩∎
+            trans (refl y) (trans y≡z z≡u) ∎)
+
+  sym-sym : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
+            sym (sym x≡y) ≡ x≡y
+  sym-sym = elim (λ {u v} u≡v → sym (sym u≡v) ≡ u≡v)
+                 (λ x → sym (sym (refl x))  ≡⟨ cong sym (sym-refl {x = x}) ⟩
+                        sym (refl x)        ≡⟨ sym-refl ⟩∎
+                        refl x              ∎)
+
+  sym-trans : ∀ {a} {A : Set a} {x y z : A}
+              (x≡y : x ≡ y) (y≡z : y ≡ z) →
+              sym (trans x≡y y≡z) ≡ trans (sym y≡z) (sym x≡y)
+  sym-trans {a} =
+    elim (λ x≡y → ∀ y≡z → sym (trans x≡y y≡z) ≡ trans (sym y≡z) (sym x≡y))
+         (λ y y≡z → sym (trans (refl y) y≡z)        ≡⟨ cong sym (trans-reflˡ y≡z) ⟩
+                    sym y≡z                         ≡⟨ sym $ trans-reflʳ (sym y≡z) ⟩
+                    trans (sym y≡z) (refl y)        ≡⟨ cong {a = a} {b = a} (trans (sym y≡z)) (sym sym-refl)  ⟩∎
+                    trans (sym y≡z) (sym (refl y))  ∎)
+
+  trans-symˡ : ∀ {a} {A : Set a} {x y : A} (p : x ≡ y) →
+               trans (sym p) p ≡ refl y
+  trans-symˡ =
+    elim (λ p → trans (sym p) p ≡ refl _)
+         (λ x → trans (sym (refl x)) (refl x)  ≡⟨ trans-reflʳ _ ⟩
+                sym (refl x)                   ≡⟨ sym-refl ⟩∎
+                refl x                         ∎)
+
+  trans-symʳ : ∀ {a} {A : Set a} {x y : A} (p : x ≡ y) →
+               trans p (sym p) ≡ refl _
+  trans-symʳ =
+    elim (λ p → trans p (sym p) ≡ refl _)
+         (λ x → trans (refl x) (sym (refl x))  ≡⟨ trans-reflˡ _ ⟩
+                sym (refl x)                   ≡⟨ sym-refl ⟩∎
+                refl x                         ∎)
+
+  cong-trans : ∀ {a b} {A : Set a} {B : Set b} {x y z : A}
+               (f : A → B) (x≡y : x ≡ y) (y≡z : y ≡ z) →
+               cong f (trans x≡y y≡z) ≡ trans (cong f x≡y) (cong f y≡z)
+  cong-trans f =
+    elim (λ x≡y → ∀ y≡z → cong f (trans x≡y y≡z) ≡
+                          trans (cong f x≡y) (cong f y≡z))
+         (λ y y≡z → cong f (trans (refl y) y≡z)           ≡⟨ cong (cong f) (trans-reflˡ _) ⟩
+                    cong f y≡z                            ≡⟨ sym $ trans-reflˡ (cong f y≡z) ⟩
+                    trans (refl (f y)) (cong f y≡z)       ≡⟨ cong₂ trans (sym (cong-refl f {x = y})) (refl (cong f y≡z)) ⟩∎
+                    trans (cong f (refl y)) (cong f y≡z)  ∎)
+
+  cong-id : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) →
+            x≡y ≡ cong id x≡y
+  cong-id = elim (λ u≡v → u≡v ≡ cong id u≡v)
+                 (λ x → refl x            ≡⟨ sym (cong-refl id {x = x}) ⟩∎
+                        cong id (refl x)  ∎)
+
+  cong-const : ∀ {a b} {A : Set a} {B : Set b} {x y : A} {z : B}
+               (x≡y : x ≡ y) →
+               cong (const z) x≡y ≡ refl z
+  cong-const {z = z} =
+    elim (λ u≡v → cong (const z) u≡v ≡ refl z)
+         (λ x → cong (const z) (refl x)  ≡⟨ cong-refl (const z) ⟩∎
+                refl z                   ∎)
+
+  cong-∘ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} {x y : A}
+           (f : B → C) (g : A → B) (x≡y : x ≡ y) →
+           cong f (cong g x≡y) ≡ cong (f ∘ g) x≡y
+  cong-∘ f g = elim (λ x≡y → cong f (cong g x≡y) ≡ cong (f ∘ g) x≡y)
+                    (λ x → cong f (cong g (refl x))  ≡⟨ cong (cong f) (cong-refl g) ⟩
+                           cong f (refl (g x))       ≡⟨ cong-refl f ⟩
+                           refl (f (g x))            ≡⟨ sym (cong-refl (f ∘ g)) ⟩∎
+                           cong (f ∘ g) (refl x)     ∎)
+
+  cong-proj₁-cong₂-, :
+    ∀ {a b} {A : Set a} {B : Set b} {x y : A} {u v : B}
+    (x≡y : x ≡ y) (u≡v : u ≡ v) →
+    cong proj₁ (cong₂ _,_ x≡y u≡v) ≡ x≡y
+  cong-proj₁-cong₂-, {x = x} {y} {u} {v} x≡y u≡v =
+    cong proj₁ (trans (cong (flip _,_ u) x≡y) (cong (_,_ y) u≡v))  ≡⟨ cong-trans proj₁ _ _ ⟩
+
+    trans (cong proj₁ (cong (flip _,_ u) x≡y))
+          (cong proj₁ (cong (_,_ y) u≡v))                          ≡⟨ cong₂ trans (cong-∘ proj₁ (flip _,_ u) x≡y) (cong-∘ proj₁ (_,_ y) u≡v) ⟩
+
+    trans (cong id x≡y) (cong (const y) u≡v)                       ≡⟨ cong₂ trans (sym $ cong-id x≡y) (cong-const u≡v) ⟩
+
+    trans x≡y (refl y)                                             ≡⟨ trans-reflʳ x≡y ⟩∎
+
+    x≡y                                                            ∎
+
+  cong-proj₂-cong₂-, :
+    ∀ {a b} {A : Set a} {B : Set b} {x y : A} {u v : B}
+    (x≡y : x ≡ y) (u≡v : u ≡ v) →
+    cong proj₂ (cong₂ _,_ x≡y u≡v) ≡ u≡v
+  cong-proj₂-cong₂-, {x = x} {y} {u} {v} x≡y u≡v =
+    cong proj₂ (trans (cong (flip _,_ u) x≡y) (cong (_,_ y) u≡v))  ≡⟨ cong-trans proj₂ _ _ ⟩
+
+    trans (cong proj₂ (cong (flip _,_ u) x≡y))
+          (cong proj₂ (cong (_,_ y) u≡v))                          ≡⟨ cong₂ trans (cong-∘ proj₂ (flip _,_ u) x≡y) (cong-∘ proj₂ (_,_ y) u≡v) ⟩
+
+    trans (cong (const u) x≡y) (cong id u≡v)                       ≡⟨ cong₂ trans (cong-const x≡y) (sym $ cong-id u≡v) ⟩
+
+    trans (refl u) u≡v                                             ≡⟨ trans-reflˡ u≡v ⟩∎
+
+    u≡v                                                            ∎
+
+  cong-sym : ∀ {a b} {A : Set a} {B : Set b} {x y : A}
+             (f : A → B) (x≡y : x ≡ y) →
+             cong f (sym x≡y) ≡ sym (cong f x≡y)
+  cong-sym f = elim (λ x≡y → cong f (sym x≡y) ≡ sym (cong f x≡y))
+                    (λ x → cong f (sym (refl x))  ≡⟨ cong (cong f) sym-refl ⟩
+                           cong f (refl x)        ≡⟨ cong-refl f ⟩
+                           refl (f x)             ≡⟨ sym sym-refl ⟩
+                           sym (refl (f x))       ≡⟨ cong sym $ sym (cong-refl f {x = x}) ⟩∎
+                           sym (cong f (refl x))  ∎)
+
+  cong₂-reflˡ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
+                (f : A → B → C) {x : A} {u v : B} {u≡v : u ≡ v} →
+                cong₂ f (refl x) u≡v ≡ cong (f x) u≡v
+  cong₂-reflˡ f {x} {u} {u≡v = u≡v} =
+    trans (cong (flip f u) (refl x)) (cong (f x) u≡v)  ≡⟨ cong₂ trans (cong-refl (flip f u)) (refl _) ⟩
+    trans (refl (f x u)) (cong (f x) u≡v)              ≡⟨ trans-reflˡ _ ⟩∎
+    cong (f x) u≡v                                     ∎
+
+  cong₂-reflʳ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
+                (f : A → B → C) {x y : A} {u : B} {x≡y : x ≡ y} →
+                cong₂ f x≡y (refl u) ≡ cong (flip f u) x≡y
+  cong₂-reflʳ f {y = y} {u} {x≡y} =
+    trans (cong (flip f u) x≡y) (cong (f y) (refl u))  ≡⟨ cong (trans _) (cong-refl (f y)) ⟩
+    trans (cong (flip f u) x≡y) (refl (f y u))         ≡⟨ trans-reflʳ _ ⟩∎
+    cong (flip f u) x≡y                                ∎
+
+  cong₂-cong-cong :
+    ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d}
+      {x₁ x₂} {eq : x₁ ≡ x₂}
+    (f : A → B) (g : A → C) (h : B → C → D) →
+    cong₂ h (cong f eq) (cong g eq) ≡
+    cong (λ x → h (f x) (g x)) eq
+  cong₂-cong-cong f g h = elim¹
+    (λ eq → cong₂ h (cong f eq) (cong g eq) ≡
+            cong (λ x → h (f x) (g x)) eq)
+    (cong₂ h (cong f (refl _)) (cong g (refl _))  ≡⟨ cong₂ (cong₂ h) (cong-refl f) (cong-refl g) ⟩
+     cong₂ h (refl _) (refl _)                    ≡⟨ cong₂-refl h ⟩
+     refl _                                       ≡⟨ sym $ cong-refl (λ x → h (f x) (g x)) ⟩∎
+     cong (λ x → h (f x) (g x)) (refl _)          ∎)
+    _
+
+  elim-∘ :
+    ∀ {a p} {A : Set a} {x y : A}
+    (P Q : ∀ {x y} → x ≡ y → Set p)
+    (f : ∀ {x y} {x≡y : x ≡ y} → P x≡y → Q x≡y)
+    (r : ∀ x → P (refl x)) {x≡y : x ≡ y} →
+    f (elim P r x≡y) ≡ elim Q (f ∘ r) x≡y
+  elim-∘ {x = x} P Q f r {x≡y} = elim¹
+    (λ x≡y → f (elim P r x≡y) ≡
+             elim Q (f ∘ r) x≡y)
+    (f (elim P r (refl x))    ≡⟨ cong f $ elim-refl P _ ⟩
+     f (r x)                  ≡⟨ sym $ elim-refl Q _ ⟩∎
+     elim Q (f ∘ r) (refl x)  ∎)
+    x≡y
+
+  elim-cong :
+    ∀ {a b p} {A : Set a} {B : Set b} {x y : A}
+    (P : B → B → Set p) (f : A → B)
+    (r : ∀ x → P x x) {x≡y : x ≡ y} →
+    elim (λ {x y} _ → P x y) r (cong f x≡y) ≡
+    elim (λ {x y} _ → P (f x) (f y)) (r ∘ f) x≡y
+  elim-cong {x = x} P f r {x≡y} = elim¹
+    (λ x≡y → elim (λ {x y} _ → P x y) r (cong f x≡y) ≡
+             elim (λ {x y} _ → P (f x) (f y)) (r ∘ f) x≡y)
+    (elim (λ {x y} _ → P x y) r (cong f (refl x))       ≡⟨ cong (elim (λ {x y} _ → P x y) _) $ cong-refl f ⟩
+     elim (λ {x y} _ → P x y) r (refl (f x))            ≡⟨ elim-refl (λ {x y} _ → P x y) _ ⟩
+     r (f x)                                            ≡⟨ sym $ elim-refl (λ {x y} _ → P (f x) (f y)) _ ⟩∎
+     elim (λ {x y} _ → P (f x) (f y)) (r ∘ f) (refl x)  ∎)
+    x≡y
hunk ./Equality.agda 730
-  abstract
-
-    subst-∘ : ∀ {a b p} {A : Set a} {B : Set b} {x y : A}
-              (P : B → Set p) (f : A → B) (x≡y : x ≡ y) {p : P (f x)} →
-              subst (P ∘ f) x≡y p ≡ subst P (cong f x≡y) p
-    subst-∘ P f _ =
-      sym $ cong (λ g → g _) $ elim-cong (λ u v → P u → P v) f _
-
-    subst-↑ : ∀ {a p ℓ} {A : Set a} {x y}
-              (P : A → Set p) {p : ↑ ℓ (P x)} {x≡y : x ≡ y} →
-              subst (↑ ℓ ∘ P) x≡y p ≡ lift (subst P x≡y (lower p))
-    subst-↑ {ℓ = ℓ} P {p} = elim¹
-      (λ x≡y → subst (↑ ℓ ∘ P) x≡y p ≡ lift (subst P x≡y (lower p)))
-      (subst (↑ ℓ ∘ P) (refl _) p         ≡⟨ subst-refl (↑ ℓ ∘ P) _ ⟩
-       p                                  ≡⟨ cong lift $ sym $ subst-refl P _ ⟩∎
-       lift (subst P (refl _) (lower p))  ∎)
-      _
-
-    -- A fusion law for subst.
-
-    subst-subst :
-      ∀ {a p} {A : Set a} (P : A → Set p)
-      {x y z : A} (x≡y : x ≡ y) (y≡z : y ≡ z) (p : P x) →
-      subst P y≡z (subst P x≡y p) ≡ subst P (trans x≡y y≡z) p
-    subst-subst P x≡y y≡z p =
-      elim (λ {x y} x≡y → ∀ {z} (y≡z : y ≡ z) p →
-              subst P y≡z (subst P x≡y p) ≡ subst P (trans x≡y y≡z) p)
-           (λ x y≡z p →
-              subst P y≡z (subst P (refl x) p)  ≡⟨ cong (subst P y≡z) $ subst-refl P p ⟩
-              subst P y≡z p                     ≡⟨ cong (λ q → subst P q p) (sym $ trans-reflˡ _) ⟩∎
-              subst P (trans (refl x) y≡z) p    ∎)
-           x≡y y≡z p
-
-    -- The combinator trans is defined in terms of subst. It could
-    -- have been defined in another way.
-
-    subst-trans :
-      ∀ {a} {A : Set a} {x y z : A} (x≡y : x ≡ y) {y≡z : y ≡ z} →
-      subst (λ x → x ≡ z) (sym x≡y) y≡z ≡ trans x≡y y≡z
-    subst-trans {y = y} {z} x≡y {y≡z} =
-      elim₁ (λ x≡y → subst (λ x → x ≡ z) (sym x≡y) y≡z ≡
-                     trans x≡y y≡z)
-            (subst (λ x → x ≡ z) (sym (refl y)) y≡z  ≡⟨ cong (λ eq → subst (λ x → x ≡ z) eq y≡z) sym-refl ⟩
-             subst (λ x → x ≡ z) (refl y) y≡z        ≡⟨ subst-refl (λ x → x ≡ z) y≡z ⟩
-             y≡z                                     ≡⟨ sym $ trans-reflˡ y≡z ⟩∎
-             trans (refl y) y≡z                      ∎)
-            x≡y
-
-    -- Substitutivity and symmetry sometimes cancel each other out.
-
-    subst-subst-sym :
-      ∀ {a p} {A : Set a} (P : A → Set p) {x y : A}
-      (x≡y : x ≡ y) (p : P y) →
-      subst P x≡y (subst P (sym x≡y) p) ≡ p
-    subst-subst-sym {A = A} P {y = y} x≡y p =
-      subst P x≡y (subst P (sym x≡y) p)  ≡⟨ subst-subst P _ _ _ ⟩
-      subst P (trans (sym x≡y) x≡y) p    ≡⟨ cong (λ q → subst P q p) (trans-symˡ x≡y) ⟩
-      subst P (refl y) p                 ≡⟨ subst-refl P p ⟩∎
-      p                                  ∎
-
-    -- Some corollaries (used in Function-universe.Π≡↔≡-↠-≡).
-
-    trans-[trans-sym] : ∀ {a} {A : Set a} {a b c : A} →
-                        (a≡b : a ≡ b) (c≡b : c ≡ b) →
-                        trans (trans a≡b (sym c≡b)) c≡b ≡ a≡b
-    trans-[trans-sym] a≡b c≡b = subst-subst-sym (_≡_ _) c≡b a≡b
-
-    trans-[trans]-sym : ∀ {a} {A : Set a} {a b c : A} →
-                        (a≡b : a ≡ b) (b≡c : b ≡ c) →
-                        trans (trans a≡b b≡c) (sym b≡c) ≡ a≡b
-    trans-[trans]-sym a≡b b≡c =
-      trans (trans a≡b b≡c)             (sym b≡c)  ≡⟨ sym $ cong (λ eq → trans (trans _ eq) (sym b≡c)) $ sym-sym _ ⟩
-      trans (trans a≡b (sym (sym b≡c))) (sym b≡c)  ≡⟨ trans-[trans-sym] _ _ ⟩∎
-      a≡b                                          ∎
-
-    -- The lemmas subst-refl and subst-const can cancel each other
-    -- out.
-
-    subst-refl-subst-const :
-      ∀ {a p} {A : Set a} {x : A} {P : Set p} {p} →
-      trans (sym $ subst-refl (λ _ → P) p) (subst-const (refl x)) ≡
-      refl p
-    subst-refl-subst-const {x = x} {P} {p} =
-      trans (sym $ subst-refl (λ _ → P) p)
-            (elim¹ (λ eq → subst (λ _ → P) eq p ≡ p)
-                   (subst-refl (λ _ → P) _) _)        ≡⟨ cong (trans _) (elim¹-refl (λ eq → subst (λ _ → P) eq p ≡ p) _) ⟩
-      trans (sym $ subst-refl (λ _ → P) p)
-            (subst-refl (λ _ → P) _)                  ≡⟨ trans-symˡ _ ⟩∎
-      refl _                                          ∎
+  subst-∘ : ∀ {a b p} {A : Set a} {B : Set b} {x y : A}
+            (P : B → Set p) (f : A → B) (x≡y : x ≡ y) {p : P (f x)} →
+            subst (P ∘ f) x≡y p ≡ subst P (cong f x≡y) p
+  subst-∘ P f _ =
+    sym $ cong (λ g → g _) $ elim-cong (λ u v → P u → P v) f _
+
+  subst-↑ : ∀ {a p ℓ} {A : Set a} {x y}
+            (P : A → Set p) {p : ↑ ℓ (P x)} {x≡y : x ≡ y} →
+            subst (↑ ℓ ∘ P) x≡y p ≡ lift (subst P x≡y (lower p))
+  subst-↑ {ℓ = ℓ} P {p} = elim¹
+    (λ x≡y → subst (↑ ℓ ∘ P) x≡y p ≡ lift (subst P x≡y (lower p)))
+    (subst (↑ ℓ ∘ P) (refl _) p         ≡⟨ subst-refl (↑ ℓ ∘ P) _ ⟩
+     p                                  ≡⟨ cong lift $ sym $ subst-refl P _ ⟩∎
+     lift (subst P (refl _) (lower p))  ∎)
+    _
+
+  -- A fusion law for subst.
+
+  subst-subst :
+    ∀ {a p} {A : Set a} (P : A → Set p)
+    {x y z : A} (x≡y : x ≡ y) (y≡z : y ≡ z) (p : P x) →
+    subst P y≡z (subst P x≡y p) ≡ subst P (trans x≡y y≡z) p
+  subst-subst P x≡y y≡z p =
+    elim (λ {x y} x≡y → ∀ {z} (y≡z : y ≡ z) p →
+            subst P y≡z (subst P x≡y p) ≡ subst P (trans x≡y y≡z) p)
+         (λ x y≡z p →
+            subst P y≡z (subst P (refl x) p)  ≡⟨ cong (subst P y≡z) $ subst-refl P p ⟩
+            subst P y≡z p                     ≡⟨ cong (λ q → subst P q p) (sym $ trans-reflˡ _) ⟩∎
+            subst P (trans (refl x) y≡z) p    ∎)
+         x≡y y≡z p
+
+  -- The combinator trans is defined in terms of subst. It could
+  -- have been defined in another way.
+
+  subst-trans :
+    ∀ {a} {A : Set a} {x y z : A} (x≡y : x ≡ y) {y≡z : y ≡ z} →
+    subst (λ x → x ≡ z) (sym x≡y) y≡z ≡ trans x≡y y≡z
+  subst-trans {y = y} {z} x≡y {y≡z} =
+    elim₁ (λ x≡y → subst (λ x → x ≡ z) (sym x≡y) y≡z ≡
+                   trans x≡y y≡z)
+          (subst (λ x → x ≡ z) (sym (refl y)) y≡z  ≡⟨ cong (λ eq → subst (λ x → x ≡ z) eq y≡z) sym-refl ⟩
+           subst (λ x → x ≡ z) (refl y) y≡z        ≡⟨ subst-refl (λ x → x ≡ z) y≡z ⟩
+           y≡z                                     ≡⟨ sym $ trans-reflˡ y≡z ⟩∎
+           trans (refl y) y≡z                      ∎)
+          x≡y
+
+  -- Substitutivity and symmetry sometimes cancel each other out.
+
+  subst-subst-sym :
+    ∀ {a p} {A : Set a} (P : A → Set p) {x y : A}
+    (x≡y : x ≡ y) (p : P y) →
+    subst P x≡y (subst P (sym x≡y) p) ≡ p
+  subst-subst-sym {A = A} P {y = y} x≡y p =
+    subst P x≡y (subst P (sym x≡y) p)  ≡⟨ subst-subst P _ _ _ ⟩
+    subst P (trans (sym x≡y) x≡y) p    ≡⟨ cong (λ q → subst P q p) (trans-symˡ x≡y) ⟩
+    subst P (refl y) p                 ≡⟨ subst-refl P p ⟩∎
+    p                                  ∎
+
+  -- Some corollaries (used in Function-universe.Π≡↔≡-↠-≡).
+
+  trans-[trans-sym] : ∀ {a} {A : Set a} {a b c : A} →
+                      (a≡b : a ≡ b) (c≡b : c ≡ b) →
+                      trans (trans a≡b (sym c≡b)) c≡b ≡ a≡b
+  trans-[trans-sym] a≡b c≡b = subst-subst-sym (_≡_ _) c≡b a≡b
+
+  trans-[trans]-sym : ∀ {a} {A : Set a} {a b c : A} →
+                      (a≡b : a ≡ b) (b≡c : b ≡ c) →
+                      trans (trans a≡b b≡c) (sym b≡c) ≡ a≡b
+  trans-[trans]-sym a≡b b≡c =
+    trans (trans a≡b b≡c)             (sym b≡c)  ≡⟨ sym $ cong (λ eq → trans (trans _ eq) (sym b≡c)) $ sym-sym _ ⟩
+    trans (trans a≡b (sym (sym b≡c))) (sym b≡c)  ≡⟨ trans-[trans-sym] _ _ ⟩∎
+    a≡b                                          ∎
+
+  -- The lemmas subst-refl and subst-const can cancel each other
+  -- out.
+
+  subst-refl-subst-const :
+    ∀ {a p} {A : Set a} {x : A} {P : Set p} {p} →
+    trans (sym $ subst-refl (λ _ → P) p) (subst-const (refl x)) ≡
+    refl p
+  subst-refl-subst-const {x = x} {P} {p} =
+    trans (sym $ subst-refl (λ _ → P) p)
+          (elim¹ (λ eq → subst (λ _ → P) eq p ≡ p)
+                 (subst-refl (λ _ → P) _) _)        ≡⟨ cong (trans _) (elim¹-refl (λ eq → subst (λ _ → P) eq p ≡ p) _) ⟩
+    trans (sym $ subst-refl (λ _ → P) p)
+          (subst-refl (λ _ → P) _)                  ≡⟨ trans-symˡ _ ⟩∎
+    refl _                                          ∎
hunk ./Equality.agda 846
-  abstract
-
-    -- "Evaluation rules" for Σ-≡,≡→≡.
-
-    Σ-≡,≡→≡-reflˡ :
-      ∀ {a b} {A : Set a} {B : A → Set b} {x y₁ y₂} →
-      (y₁≡y₂ : subst B (refl x) y₁ ≡ y₂) →
-      Σ-≡,≡→≡ {B = B} (refl x) y₁≡y₂ ≡
-      cong (_,_ x) (trans (sym $ subst-refl B y₁) y₁≡y₂)
-    Σ-≡,≡→≡-reflˡ {B = B} y₁≡y₂ =
-      cong (λ f → f y₁≡y₂) $
-        elim-refl (λ {x₁ y₁} (p : x₁ ≡ y₁) → ∀ {x₂ y₂} →
-                     subst B p x₂ ≡ y₂ → (x₁ , x₂) ≡ (y₁ , y₂))
-                  _
-
-    Σ-≡,≡→≡-refl-refl :
-      ∀ {a b} {A : Set a} {B : A → Set b} {x y} →
-      Σ-≡,≡→≡ {B = B} (refl x) (refl (subst B (refl x) y)) ≡
-      cong (_,_ x) (sym (subst-refl B y))
-    Σ-≡,≡→≡-refl-refl {B = B} {x} {y} =
-      Σ-≡,≡→≡ (refl x) (refl _)                             ≡⟨ Σ-≡,≡→≡-reflˡ (refl _) ⟩
-      cong (_,_ x) (trans (sym $ subst-refl B y) (refl _))  ≡⟨ cong (cong (_,_ x)) (trans-reflʳ _) ⟩∎
-      cong (_,_ x) (sym (subst-refl B y))                   ∎
-
-    -- "Evaluation rule" for Σ-≡,≡←≡.
-
-    Σ-≡,≡←≡-refl :
-      ∀ {a b} {A : Set a} {B : A → Set b} {p : Σ A B} →
-      Σ-≡,≡←≡ (refl p) ≡ (refl (proj₁ p) , subst-refl B (proj₂ p))
-    Σ-≡,≡←≡-refl {A = A} {B} = elim-refl
-      (λ {p₁ p₂ : Σ A B} _ →
-         ∃ λ (p : proj₁ p₁ ≡ proj₁ p₂) →
-           subst B p (proj₂ p₁) ≡ proj₂ p₂)
-      _
-
-    -- Proof simplification rules for Σ-≡,≡→≡.
-
-    proj₁-Σ-≡,≡→≡ :
-      ∀ {a b} {A : Set a} {B : A → Set b} {x₁ x₂ y₁ y₂}
-      (x₁≡x₂ : x₁ ≡ x₂) (y₁≡y₂ : subst B x₁≡x₂ y₁ ≡ y₂) →
-      cong proj₁ (Σ-≡,≡→≡ {B = B} x₁≡x₂ y₁≡y₂) ≡ x₁≡x₂
-    proj₁-Σ-≡,≡→≡ {B = B} {y₁ = y₁} x₁≡x₂ y₁≡y₂ = elim¹
-      (λ x₁≡x₂ → ∀ {y₂} (y₁≡y₂ : subst B x₁≡x₂ y₁ ≡ y₂) →
-         cong proj₁ (Σ-≡,≡→≡ x₁≡x₂ y₁≡y₂) ≡ x₁≡x₂)
-      (λ y₁≡y₂ →
-         cong proj₁ (Σ-≡,≡→≡ (refl _) y₁≡y₂)                              ≡⟨ cong (cong proj₁) $ Σ-≡,≡→≡-reflˡ y₁≡y₂ ⟩
-         cong proj₁ (cong (_,_ _) (trans (sym $ subst-refl B y₁) y₁≡y₂))  ≡⟨ cong-∘ proj₁ (_,_ _) _ ⟩
-         cong (const _) (trans (sym $ subst-refl B y₁) y₁≡y₂)             ≡⟨ cong-const _ ⟩∎
-         refl _                                                           ∎)
-      x₁≡x₂ y₁≡y₂
-
-    Σ-≡,≡→≡-subst-const :
-      ∀ {a b} {A : Set a} {B : Set b} {p₁ p₂ : A × B} →
-      (p : proj₁ p₁ ≡ proj₁ p₂) (q : proj₂ p₁ ≡ proj₂ p₂) →
-      Σ-≡,≡→≡ p (trans (subst-const p) q) ≡ cong₂ _,_ p q
-    Σ-≡,≡→≡-subst-const {B = B} {_ , y₁} {_ , y₂} p q = elim
-      (λ {x₁ y₁} (p : x₁ ≡ y₁) →
-         Σ-≡,≡→≡ p (trans (subst-const _) q) ≡ cong₂ _,_ p q)
-      (λ x →
-         let lemma =
-               trans (sym $ subst-refl (λ _ → B) y₁)
-                     (trans (subst-const _) q)               ≡⟨ sym $ trans-assoc _ _ _ ⟩
-               trans (trans (sym $ subst-refl (λ _ → B) y₁)
-                            (subst-const _))
-                     q                                       ≡⟨ cong₂ trans subst-refl-subst-const (refl _) ⟩
-               trans (refl y₁) q                             ≡⟨ trans-reflˡ _ ⟩∎
-               q                                             ∎ in
-
-         Σ-≡,≡→≡ (refl x) (trans (subst-const _) q)           ≡⟨ Σ-≡,≡→≡-reflˡ _ ⟩
-         cong (_,_ x) (trans (sym $ subst-refl (λ _ → B) y₁)
-                             (trans (subst-const _) q))       ≡⟨ cong (cong (_,_ x)) lemma ⟩
-         cong (_,_ x) q                                       ≡⟨ sym $ cong₂-reflˡ _,_ ⟩∎
-         cong₂ _,_ (refl x) q                                 ∎)
-      p
-
-    -- Proof simplification rule for Σ-≡,≡←≡.
-
-    proj₁-Σ-≡,≡←≡ :
-      ∀ {a b} {A : Set a} {B : A → Set b} {p₁ p₂ : Σ A B}
-      (p₁≡p₂ : p₁ ≡ p₂) →
-      proj₁ (Σ-≡,≡←≡ p₁≡p₂) ≡ cong proj₁ p₁≡p₂
-    proj₁-Σ-≡,≡←≡ = elim
-      (λ p₁≡p₂ → proj₁ (Σ-≡,≡←≡ p₁≡p₂) ≡ cong proj₁ p₁≡p₂)
-      (λ p →
-         proj₁ (Σ-≡,≡←≡ (refl p))  ≡⟨ cong proj₁ $ Σ-≡,≡←≡-refl ⟩
-         refl (proj₁ p)            ≡⟨ sym $ cong-refl proj₁ ⟩∎
-         cong proj₁ (refl p)       ∎)
+  -- "Evaluation rules" for Σ-≡,≡→≡.
+
+  Σ-≡,≡→≡-reflˡ :
+    ∀ {a b} {A : Set a} {B : A → Set b} {x y₁ y₂} →
+    (y₁≡y₂ : subst B (refl x) y₁ ≡ y₂) →
+    Σ-≡,≡→≡ {B = B} (refl x) y₁≡y₂ ≡
+    cong (_,_ x) (trans (sym $ subst-refl B y₁) y₁≡y₂)
+  Σ-≡,≡→≡-reflˡ {B = B} y₁≡y₂ =
+    cong (λ f → f y₁≡y₂) $
+      elim-refl (λ {x₁ y₁} (p : x₁ ≡ y₁) → ∀ {x₂ y₂} →
+                   subst B p x₂ ≡ y₂ → (x₁ , x₂) ≡ (y₁ , y₂))
+                _
+
+  Σ-≡,≡→≡-refl-refl :
+    ∀ {a b} {A : Set a} {B : A → Set b} {x y} →
+    Σ-≡,≡→≡ {B = B} (refl x) (refl (subst B (refl x) y)) ≡
+    cong (_,_ x) (sym (subst-refl B y))
+  Σ-≡,≡→≡-refl-refl {B = B} {x} {y} =
+    Σ-≡,≡→≡ (refl x) (refl _)                             ≡⟨ Σ-≡,≡→≡-reflˡ (refl _) ⟩
+    cong (_,_ x) (trans (sym $ subst-refl B y) (refl _))  ≡⟨ cong (cong (_,_ x)) (trans-reflʳ _) ⟩∎
+    cong (_,_ x) (sym (subst-refl B y))                   ∎
+
+  -- "Evaluation rule" for Σ-≡,≡←≡.
+
+  Σ-≡,≡←≡-refl :
+    ∀ {a b} {A : Set a} {B : A → Set b} {p : Σ A B} →
+    Σ-≡,≡←≡ (refl p) ≡ (refl (proj₁ p) , subst-refl B (proj₂ p))
+  Σ-≡,≡←≡-refl {A = A} {B} = elim-refl
+    (λ {p₁ p₂ : Σ A B} _ →
+       ∃ λ (p : proj₁ p₁ ≡ proj₁ p₂) →
+         subst B p (proj₂ p₁) ≡ proj₂ p₂)
+    _
+
+  -- Proof simplification rules for Σ-≡,≡→≡.
+
+  proj₁-Σ-≡,≡→≡ :
+    ∀ {a b} {A : Set a} {B : A → Set b} {x₁ x₂ y₁ y₂}
+    (x₁≡x₂ : x₁ ≡ x₂) (y₁≡y₂ : subst B x₁≡x₂ y₁ ≡ y₂) →
+    cong proj₁ (Σ-≡,≡→≡ {B = B} x₁≡x₂ y₁≡y₂) ≡ x₁≡x₂
+  proj₁-Σ-≡,≡→≡ {B = B} {y₁ = y₁} x₁≡x₂ y₁≡y₂ = elim¹
+    (λ x₁≡x₂ → ∀ {y₂} (y₁≡y₂ : subst B x₁≡x₂ y₁ ≡ y₂) →
+       cong proj₁ (Σ-≡,≡→≡ x₁≡x₂ y₁≡y₂) ≡ x₁≡x₂)
+    (λ y₁≡y₂ →
+       cong proj₁ (Σ-≡,≡→≡ (refl _) y₁≡y₂)                              ≡⟨ cong (cong proj₁) $ Σ-≡,≡→≡-reflˡ y₁≡y₂ ⟩
+       cong proj₁ (cong (_,_ _) (trans (sym $ subst-refl B y₁) y₁≡y₂))  ≡⟨ cong-∘ proj₁ (_,_ _) _ ⟩
+       cong (const _) (trans (sym $ subst-refl B y₁) y₁≡y₂)             ≡⟨ cong-const _ ⟩∎
+       refl _                                                           ∎)
+    x₁≡x₂ y₁≡y₂
+
+  Σ-≡,≡→≡-subst-const :
+    ∀ {a b} {A : Set a} {B : Set b} {p₁ p₂ : A × B} →
+    (p : proj₁ p₁ ≡ proj₁ p₂) (q : proj₂ p₁ ≡ proj₂ p₂) →
+    Σ-≡,≡→≡ p (trans (subst-const p) q) ≡ cong₂ _,_ p q
+  Σ-≡,≡→≡-subst-const {B = B} {_ , y₁} {_ , y₂} p q = elim
+    (λ {x₁ y₁} (p : x₁ ≡ y₁) →
+       Σ-≡,≡→≡ p (trans (subst-const _) q) ≡ cong₂ _,_ p q)
+    (λ x →
+       let lemma =
+             trans (sym $ subst-refl (λ _ → B) y₁)
+                   (trans (subst-const _) q)               ≡⟨ sym $ trans-assoc _ _ _ ⟩
+             trans (trans (sym $ subst-refl (λ _ → B) y₁)
+                          (subst-const _))
+                   q                                       ≡⟨ cong₂ trans subst-refl-subst-const (refl _) ⟩
+             trans (refl y₁) q                             ≡⟨ trans-reflˡ _ ⟩∎
+             q                                             ∎ in
+
+       Σ-≡,≡→≡ (refl x) (trans (subst-const _) q)           ≡⟨ Σ-≡,≡→≡-reflˡ _ ⟩
+       cong (_,_ x) (trans (sym $ subst-refl (λ _ → B) y₁)
+                           (trans (subst-const _) q))       ≡⟨ cong (cong (_,_ x)) lemma ⟩
+       cong (_,_ x) q                                       ≡⟨ sym $ cong₂-reflˡ _,_ ⟩∎
+       cong₂ _,_ (refl x) q                                 ∎)
+    p
+
+  -- Proof simplification rule for Σ-≡,≡←≡.
+
+  proj₁-Σ-≡,≡←≡ :
+    ∀ {a b} {A : Set a} {B : A → Set b} {p₁ p₂ : Σ A B}
+    (p₁≡p₂ : p₁ ≡ p₂) →
+    proj₁ (Σ-≡,≡←≡ p₁≡p₂) ≡ cong proj₁ p₁≡p₂
+  proj₁-Σ-≡,≡←≡ = elim
+    (λ p₁≡p₂ → proj₁ (Σ-≡,≡←≡ p₁≡p₂) ≡ cong proj₁ p₁≡p₂)
+    (λ p →
+       proj₁ (Σ-≡,≡←≡ (refl p))  ≡⟨ cong proj₁ $ Σ-≡,≡←≡-refl ⟩
+       refl (proj₁ p)            ≡⟨ sym $ cong-refl proj₁ ⟩∎
+       cong proj₁ (refl p)       ∎)
hunk ./Equality.agda 940
-  abstract
-
-    -- "Evaluation rule" for subst₂.
-
-    subst₂-refl-refl :
-      ∀ {a b p} {A : Set a} {B : A → Set b}
-      (P : Σ A B → Set p) {x y p} →
-      subst₂ P (refl _) (refl _) p ≡
-      subst (curry P x) (sym $ subst-refl B y) p
-    subst₂-refl-refl {B = B} P {x} {y} {p} =
-      subst P (Σ-≡,≡→≡ (refl x) (refl _)) p            ≡⟨ cong (λ eq₁ → subst P eq₁ p) Σ-≡,≡→≡-refl-refl ⟩
-      subst P (cong (_,_ x) (sym (subst-refl B y))) p  ≡⟨ sym $ subst-∘ P (_,_ x) _ ⟩∎
-      subst (curry P x) (sym $ subst-refl B y) p       ∎
-
-    -- The subst function can be "pushed" inside pairs.
-
-    push-subst-pair :
-      ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
-      (B : A → Set b) (C : Σ A B → Set c) {p} →
-      subst (λ x → Σ (B x) (curry C x)) y≡z p ≡
-      (subst B y≡z (proj₁ p) , subst₂ C y≡z (refl _) (proj₂ p))
-    push-subst-pair {y≡z = y≡z} B C {p} = elim¹
-      (λ y≡z →
-         subst (λ x → Σ (B x) (curry C x)) y≡z p ≡
-         (subst B y≡z (proj₁ p) , subst₂ C y≡z (refl _) (proj₂ p)))
-      (subst (λ x → Σ (B x) (curry C x)) (refl _) p  ≡⟨ subst-refl (λ x → Σ (B x) (curry C x)) _ ⟩
-       p                                             ≡⟨ Σ-≡,≡→≡ (sym (subst-refl B _)) (sym (subst₂-refl-refl C)) ⟩∎
-       (subst B (refl _) (proj₁ p) ,
-        subst₂ C (refl _) (refl _) (proj₂ p))        ∎)
-      y≡z
-
-    -- A corollary of push-subst-pair.
-
-    push-subst-pair′ :
-      ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
-      (B : A → Set b) (C : Σ A B → Set c) {p p₁} →
-      (p₁≡p₁ : subst B y≡z (proj₁ p) ≡ p₁) →
-      subst (λ x → Σ (B x) (curry C x)) y≡z p ≡
-      (p₁ , subst₂ C y≡z p₁≡p₁ (proj₂ p))
-    push-subst-pair′ {y≡z = y≡z} B C {p} =
-      elim¹ (λ {p₁} p₁≡p₁ →
-               subst (λ x → Σ (B x) (curry C x)) y≡z p ≡
-               (p₁ , subst₂ C y≡z p₁≡p₁ (proj₂ p)))
-            (push-subst-pair B C)
-
-    -- A proof simplification rule for subst₂.
-
-    subst₂-proj₁ :
-      ∀ {a b p} {A : Set a} {B : A → Set b} {x₁ x₂ y₁ y₂}
-        {x₁≡x₂ : x₁ ≡ x₂} {y₁≡y₂ : subst B x₁≡x₂ y₁ ≡ y₂}
-      (P : A → Set p) {p} →
-      subst₂ {B = B} (P ∘ proj₁) x₁≡x₂ y₁≡y₂ p ≡ subst P x₁≡x₂ p
-    subst₂-proj₁ {x₁≡x₂ = x₁≡x₂} {y₁≡y₂} P {p} =
-      subst₂ (P ∘ proj₁) x₁≡x₂ y₁≡y₂ p              ≡⟨ subst-∘ P proj₁ _ ⟩
-      subst P (cong proj₁ (Σ-≡,≡→≡ x₁≡x₂ y₁≡y₂)) p  ≡⟨ cong (λ eq → subst P eq p) (proj₁-Σ-≡,≡→≡ _ _) ⟩∎
-      subst P x₁≡x₂ p                               ∎
-
-    -- The subst function can be "pushed" inside non-dependent pairs.
-
-    push-subst-, :
-      ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
-      (B : A → Set b) (C : A → Set c) {p} →
-      subst (λ x → B x × C x) y≡z p ≡
-      (subst B y≡z (proj₁ p) , subst C y≡z (proj₂ p))
-    push-subst-, {y≡z = y≡z} B C {x , y} =
-      subst (λ x → B x × C x) y≡z (x , y)                           ≡⟨ push-subst-pair B (C ∘ proj₁) ⟩
-      (subst B y≡z x , subst (C ∘ proj₁) (Σ-≡,≡→≡ y≡z (refl _)) y)  ≡⟨ cong (_,_ _) $ subst₂-proj₁ C ⟩∎
-      (subst B y≡z x , subst C y≡z y)                               ∎
-
-    -- The subst function can be "pushed" inside inj₁ and inj₂.
-
-    push-subst-inj₁ :
-      ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
-      (B : A → Set b) (C : A → Set c) {x} →
-      subst (λ x → B x ⊎ C x) y≡z (inj₁ x) ≡ inj₁ (subst B y≡z x)
-    push-subst-inj₁ {y≡z = y≡z} B C {x} = elim¹
-      (λ y≡z → subst (λ x → B x ⊎ C x) y≡z (inj₁ x) ≡
-               inj₁ (subst B y≡z x))
-      (subst (λ x → B x ⊎ C x) (refl _) (inj₁ x)  ≡⟨ subst-refl (λ x → B x ⊎ C x) _ ⟩
-       inj₁ x                                     ≡⟨ cong inj₁ $ sym $ subst-refl B _ ⟩∎
-       inj₁ (subst B (refl _) x)                  ∎)
-      y≡z
-
-    push-subst-inj₂ :
-      ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
-      (B : A → Set b) (C : A → Set c) {x} →
-      subst (λ x → B x ⊎ C x) y≡z (inj₂ x) ≡ inj₂ (subst C y≡z x)
-    push-subst-inj₂ {y≡z = y≡z} B C {x} = elim¹
-      (λ y≡z → subst (λ x → B x ⊎ C x) y≡z (inj₂ x) ≡
-               inj₂ (subst C y≡z x))
-      (subst (λ x → B x ⊎ C x) (refl _) (inj₂ x)  ≡⟨ subst-refl (λ x → B x ⊎ C x) _ ⟩
-       inj₂ x                                     ≡⟨ cong inj₂ $ sym $ subst-refl C _ ⟩∎
-       inj₂ (subst C (refl _) x)                  ∎)
-      y≡z
-
-    -- The subst function can be "pushed" inside applications.
-
-    push-subst-application :
-      ∀ {a b c} {A : Set a} {x : A} {B : Set b} {y₁ y₂ : B}
-      (y₁≡y₂ : y₁ ≡ y₂) (C : A → B → Set c) {f : (x : A) → C x y₁} →
-      subst (C x) y₁≡y₂ (f x) ≡
-      subst (λ y → (x : A) → C x y) y₁≡y₂ f x
-    push-subst-application {x = x} y₁≡y₂ C {f} = elim¹
-      (λ y₁≡y₂ → subst (C x) y₁≡y₂ (f x) ≡
-                 subst (λ y → ∀ x → C x y) y₁≡y₂ f x)
-      (subst (C x) (refl _) (f x)              ≡⟨ subst-refl (C x) _ ⟩
-       f x                                     ≡⟨ cong (λ g → g x) $ sym $ subst-refl (λ y → ∀ x → C x y) _ ⟩∎
-       subst (λ y → ∀ x → C x y) (refl _) f x  ∎)
+  -- "Evaluation rule" for subst₂.
+
+  subst₂-refl-refl :
+    ∀ {a b p} {A : Set a} {B : A → Set b}
+    (P : Σ A B → Set p) {x y p} →
+    subst₂ P (refl _) (refl _) p ≡
+    subst (curry P x) (sym $ subst-refl B y) p
+  subst₂-refl-refl {B = B} P {x} {y} {p} =
+    subst P (Σ-≡,≡→≡ (refl x) (refl _)) p            ≡⟨ cong (λ eq₁ → subst P eq₁ p) Σ-≡,≡→≡-refl-refl ⟩
+    subst P (cong (_,_ x) (sym (subst-refl B y))) p  ≡⟨ sym $ subst-∘ P (_,_ x) _ ⟩∎
+    subst (curry P x) (sym $ subst-refl B y) p       ∎
+
+  -- The subst function can be "pushed" inside pairs.
+
+  push-subst-pair :
+    ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
+    (B : A → Set b) (C : Σ A B → Set c) {p} →
+    subst (λ x → Σ (B x) (curry C x)) y≡z p ≡
+    (subst B y≡z (proj₁ p) , subst₂ C y≡z (refl _) (proj₂ p))
+  push-subst-pair {y≡z = y≡z} B C {p} = elim¹
+    (λ y≡z →
+       subst (λ x → Σ (B x) (curry C x)) y≡z p ≡
+       (subst B y≡z (proj₁ p) , subst₂ C y≡z (refl _) (proj₂ p)))
+    (subst (λ x → Σ (B x) (curry C x)) (refl _) p  ≡⟨ subst-refl (λ x → Σ (B x) (curry C x)) _ ⟩
+     p                                             ≡⟨ Σ-≡,≡→≡ (sym (subst-refl B _)) (sym (subst₂-refl-refl C)) ⟩∎
+     (subst B (refl _) (proj₁ p) ,
+      subst₂ C (refl _) (refl _) (proj₂ p))        ∎)
+    y≡z
+
+  -- A corollary of push-subst-pair.
+
+  push-subst-pair′ :
+    ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
+    (B : A → Set b) (C : Σ A B → Set c) {p p₁} →
+    (p₁≡p₁ : subst B y≡z (proj₁ p) ≡ p₁) →
+    subst (λ x → Σ (B x) (curry C x)) y≡z p ≡
+    (p₁ , subst₂ C y≡z p₁≡p₁ (proj₂ p))
+  push-subst-pair′ {y≡z = y≡z} B C {p} =
+    elim¹ (λ {p₁} p₁≡p₁ →
+             subst (λ x → Σ (B x) (curry C x)) y≡z p ≡
+             (p₁ , subst₂ C y≡z p₁≡p₁ (proj₂ p)))
+          (push-subst-pair B C)
+
+  -- A proof simplification rule for subst₂.
+
+  subst₂-proj₁ :
+    ∀ {a b p} {A : Set a} {B : A → Set b} {x₁ x₂ y₁ y₂}
+      {x₁≡x₂ : x₁ ≡ x₂} {y₁≡y₂ : subst B x₁≡x₂ y₁ ≡ y₂}
+    (P : A → Set p) {p} →
+    subst₂ {B = B} (P ∘ proj₁) x₁≡x₂ y₁≡y₂ p ≡ subst P x₁≡x₂ p
+  subst₂-proj₁ {x₁≡x₂ = x₁≡x₂} {y₁≡y₂} P {p} =
+    subst₂ (P ∘ proj₁) x₁≡x₂ y₁≡y₂ p              ≡⟨ subst-∘ P proj₁ _ ⟩
+    subst P (cong proj₁ (Σ-≡,≡→≡ x₁≡x₂ y₁≡y₂)) p  ≡⟨ cong (λ eq → subst P eq p) (proj₁-Σ-≡,≡→≡ _ _) ⟩∎
+    subst P x₁≡x₂ p                               ∎
+
+  -- The subst function can be "pushed" inside non-dependent pairs.
+
+  push-subst-, :
+    ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
+    (B : A → Set b) (C : A → Set c) {p} →
+    subst (λ x → B x × C x) y≡z p ≡
+    (subst B y≡z (proj₁ p) , subst C y≡z (proj₂ p))
+  push-subst-, {y≡z = y≡z} B C {x , y} =
+    subst (λ x → B x × C x) y≡z (x , y)                           ≡⟨ push-subst-pair B (C ∘ proj₁) ⟩
+    (subst B y≡z x , subst (C ∘ proj₁) (Σ-≡,≡→≡ y≡z (refl _)) y)  ≡⟨ cong (_,_ _) $ subst₂-proj₁ C ⟩∎
+    (subst B y≡z x , subst C y≡z y)                               ∎
+
+  -- The subst function can be "pushed" inside inj₁ and inj₂.
+
+  push-subst-inj₁ :
+    ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
+    (B : A → Set b) (C : A → Set c) {x} →
+    subst (λ x → B x ⊎ C x) y≡z (inj₁ x) ≡ inj₁ (subst B y≡z x)
+  push-subst-inj₁ {y≡z = y≡z} B C {x} = elim¹
+    (λ y≡z → subst (λ x → B x ⊎ C x) y≡z (inj₁ x) ≡
+             inj₁ (subst B y≡z x))
+    (subst (λ x → B x ⊎ C x) (refl _) (inj₁ x)  ≡⟨ subst-refl (λ x → B x ⊎ C x) _ ⟩
+     inj₁ x                                     ≡⟨ cong inj₁ $ sym $ subst-refl B _ ⟩∎
+     inj₁ (subst B (refl _) x)                  ∎)
+    y≡z
+
+  push-subst-inj₂ :
+    ∀ {a b c} {A : Set a} {y z : A} {y≡z : y ≡ z}
+    (B : A → Set b) (C : A → Set c) {x} →
+    subst (λ x → B x ⊎ C x) y≡z (inj₂ x) ≡ inj₂ (subst C y≡z x)
+  push-subst-inj₂ {y≡z = y≡z} B C {x} = elim¹
+    (λ y≡z → subst (λ x → B x ⊎ C x) y≡z (inj₂ x) ≡
+             inj₂ (subst C y≡z x))
+    (subst (λ x → B x ⊎ C x) (refl _) (inj₂ x)  ≡⟨ subst-refl (λ x → B x ⊎ C x) _ ⟩
+     inj₂ x                                     ≡⟨ cong inj₂ $ sym $ subst-refl C _ ⟩∎
+     inj₂ (subst C (refl _) x)                  ∎)
+    y≡z
+
+  -- The subst function can be "pushed" inside applications.
+
+  push-subst-application :
+    ∀ {a b c} {A : Set a} {x : A} {B : Set b} {y₁ y₂ : B}
+    (y₁≡y₂ : y₁ ≡ y₂) (C : A → B → Set c) {f : (x : A) → C x y₁} →
+    subst (C x) y₁≡y₂ (f x) ≡
+    subst (λ y → (x : A) → C x y) y₁≡y₂ f x
+  push-subst-application {x = x} y₁≡y₂ C {f} = elim¹
+    (λ y₁≡y₂ → subst (C x) y₁≡y₂ (f x) ≡
+               subst (λ y → ∀ x → C x y) y₁≡y₂ f x)
+    (subst (C x) (refl _) (f x)              ≡⟨ subst-refl (C x) _ ⟩
+     f x                                     ≡⟨ cong (λ g → g x) $ sym $ subst-refl (λ y → ∀ x → C x y) _ ⟩∎
+     subst (λ y → ∀ x → C x y) (refl _) f x  ∎)
+    y₁≡y₂
+
+  push-subst-implicit-application :
+    ∀ {a b c} {A : Set a} {x : A} {B : Set b} {y₁ y₂ : B}
+    (y₁≡y₂ : y₁ ≡ y₂) (C : A → B → Set c) {f : {x : A} → C x y₁} →
+    subst (C x) y₁≡y₂ (f {x = x}) ≡
+    subst (λ y → {x : A} → C x y) y₁≡y₂ f {x = x}
+  push-subst-implicit-application {x = x} y₁≡y₂ C {f} =
+    elim¹
+      (λ y₁≡y₂ → subst (C x) y₁≡y₂ f ≡
+                 subst (λ y → ∀ {x} → C x y) y₁≡y₂ f)
+      (subst (C x) (refl _) f                  ≡⟨ subst-refl (C x) _ ⟩
+       f                                       ≡⟨ cong (λ g → g {x = x}) $ sym $ subst-refl (λ y → ∀ {x} → C x y) _ ⟩∎
+       subst (λ y → ∀ {x} → C x y) (refl _) f  ∎)
hunk ./Equality.agda 1062
-    push-subst-implicit-application :
-      ∀ {a b c} {A : Set a} {x : A} {B : Set b} {y₁ y₂ : B}
-      (y₁≡y₂ : y₁ ≡ y₂) (C : A → B → Set c) {f : {x : A} → C x y₁} →
-      subst (C x) y₁≡y₂ (f {x = x}) ≡
-      subst (λ y → {x : A} → C x y) y₁≡y₂ f {x = x}
-    push-subst-implicit-application {x = x} y₁≡y₂ C {f} =
-      elim¹
-        (λ y₁≡y₂ → subst (C x) y₁≡y₂ f ≡
-                   subst (λ y → ∀ {x} → C x y) y₁≡y₂ f)
-        (subst (C x) (refl _) f                  ≡⟨ subst-refl (C x) _ ⟩
-         f                                       ≡⟨ cong (λ g → g {x = x}) $ sym $ subst-refl (λ y → ∀ {x} → C x y) _ ⟩∎
-         subst (λ y → ∀ {x} → C x y) (refl _) f  ∎)
-        y₁≡y₂
-
-    -- An application of subst to the right kind of function
-    -- application is equal to a function application.
-
-    subst-application :
-      ∀ {a p} {A : Set a} {P : A → Set p}
-      (f : ∀ x → P x) {x y} (x≡y : x ≡ y) →
-      subst P x≡y (f x) ≡ f y
-    subst-application {P = P} f =
-      elim (λ {x y} x≡y → subst P x≡y (f x) ≡ f y)
-           (subst-refl P ∘ f)
-
-    -- If f z evaluates to z for a decidable set of values which
-    -- includes x and y, do we have
-    --
-    --   cong f x≡y ≡ x≡y
-    --
-    -- for any x≡y : x ≡ y? The equation above is not well-typed if f
-    -- is a variable, but the approximation below can be proved.
-
-    cong-roughly-id : ∀ {a} {A : Set a} (f : A → A) (p : A → Bool) {x y : A}
-                      (x≡y : x ≡ y) (px : T (p x)) (py : T (p y))
-                      (f≡id : ∀ z → T (p z) → f z ≡ z) →
-                      cong f x≡y ≡
-                      trans (f≡id x px) (trans x≡y $ sym (f≡id y py))
-    cong-roughly-id {A = A} f p =
-      elim (λ {x y} x≡y →
-              (px : T (p x)) (py : T (p y))
-              (f≡id : ∀ z → T (p z) → f z ≡ z) →
-              cong f x≡y ≡
-              trans (f≡id x px) (trans x≡y $ sym (f≡id y py)))
-           (λ x px px′ f≡id → helper x (p x) px px′ (f≡id x))
-      where
-      helper :
-        (x : A) (b : Bool) (px px′ : T b)
-        (f≡id : T b → f x ≡ x) →
-        cong f (refl x) ≡
-        trans (f≡id px) (trans (refl x) $ sym (f≡id px′))
-      helper x false px _ f≡id = ⊥-elim px
-      helper x true  _  _ f≡id =
-        cong f (refl x)                                 ≡⟨ cong-refl f ⟩
-        refl (f x)                                      ≡⟨ sym $ trans-symʳ _ ⟩
-        trans (f≡id _) (sym (f≡id _))                   ≡⟨ cong (trans (f≡id _)) $ sym $ trans-reflˡ _ ⟩∎
-        trans (f≡id _) (trans (refl x) $ sym (f≡id _))  ∎
-
-    -- The following lemma is Proposition 2 from "Generalizations of
-    -- Hedberg's Theorem" by Kraus, Escardó, Coquand and Altenkirch.
-
-    subst-in-terms-of-trans-and-cong :
-      ∀ {a b} {A : Set a} {B : Set b} {f g : A → B} {x y}
-        {x≡y : x ≡ y} {fx≡gx : f x ≡ g x} →
-      subst (λ z → f z ≡ g z) x≡y fx≡gx ≡
-      trans (sym (cong f x≡y)) (trans fx≡gx (cong g x≡y))
-    subst-in-terms-of-trans-and-cong {f = f} {g} = elim
-      (λ {x y} x≡y →
-           (fx≡gx : f x ≡ g x) →
-           subst (λ z → f z ≡ g z) x≡y fx≡gx ≡
-           trans (sym (cong f x≡y)) (trans fx≡gx (cong g x≡y)))
-      (λ x fx≡gx →
-           subst (λ z → f z ≡ g z) (refl x) fx≡gx                         ≡⟨ subst-refl _ _ ⟩
-           fx≡gx                                                          ≡⟨ sym $ trans-reflˡ _ ⟩
-           trans (refl (f x)) fx≡gx                                       ≡⟨ sym $ cong₂ trans sym-refl (trans-reflʳ _)  ⟩
-           trans (sym (refl (f x))) (trans fx≡gx (refl (g x)))            ≡⟨ sym $ cong₂ (λ p q → trans (sym p) (trans _ q))
-                                                                                         (cong-refl _) (cong-refl _) ⟩∎
-           trans (sym (cong f (refl x))) (trans fx≡gx (cong g (refl x)))  ∎ )
-      _
-      _
-
-    -- The following lemma is basically Theorem 2.11.5 from the HoTT
-    -- book (the book's lemma gives an equivalence between equality
-    -- types, rather than an equality between equality types).
-
-    [subst≡]≡[trans≡trans] :
-      ∀ {a} {A : Set a} {x y : A} {p : x ≡ y} {q : x ≡ x} {r : y ≡ y} →
-      (subst (λ z → z ≡ z) p q ≡ r)
-        ≡
-      (trans q p ≡ trans p r)
-    [subst≡]≡[trans≡trans] {p = p} {q} {r} = elim
-      (λ {x y} p → {q : x ≡ x} {r : y ≡ y} →
-                   (subst (λ z → z ≡ z) p q ≡ r)
-                     ≡
-                   (trans q p ≡ trans p r))
-      (λ x {q r} →
-         subst (λ z → z ≡ z) (refl x) q ≡ r   ≡⟨ cong (_≡ _) (subst-refl (λ z → z ≡ z) _) ⟩
-         q ≡ r                                ≡⟨ sym $ cong₂ _≡_ (trans-reflʳ _) (trans-reflˡ _) ⟩∎
-         trans q (refl x) ≡ trans (refl x) r  ∎)
-      p
+  -- An application of subst to the right kind of function
+  -- application is equal to a function application.
+
+  subst-application :
+    ∀ {a p} {A : Set a} {P : A → Set p}
+    (f : ∀ x → P x) {x y} (x≡y : x ≡ y) →
+    subst P x≡y (f x) ≡ f y
+  subst-application {P = P} f =
+    elim (λ {x y} x≡y → subst P x≡y (f x) ≡ f y)
+         (subst-refl P ∘ f)
+
+  -- If f z evaluates to z for a decidable set of values which
+  -- includes x and y, do we have
+  --
+  --   cong f x≡y ≡ x≡y
+  --
+  -- for any x≡y : x ≡ y? The equation above is not well-typed if f
+  -- is a variable, but the approximation below can be proved.
+
+  cong-roughly-id : ∀ {a} {A : Set a} (f : A → A) (p : A → Bool) {x y : A}
+                    (x≡y : x ≡ y) (px : T (p x)) (py : T (p y))
+                    (f≡id : ∀ z → T (p z) → f z ≡ z) →
+                    cong f x≡y ≡
+                    trans (f≡id x px) (trans x≡y $ sym (f≡id y py))
+  cong-roughly-id {A = A} f p =
+    elim (λ {x y} x≡y →
+            (px : T (p x)) (py : T (p y))
+            (f≡id : ∀ z → T (p z) → f z ≡ z) →
+            cong f x≡y ≡
+            trans (f≡id x px) (trans x≡y $ sym (f≡id y py)))
+         (λ x px px′ f≡id → helper x (p x) px px′ (f≡id x))
+    where
+    helper :
+      (x : A) (b : Bool) (px px′ : T b)
+      (f≡id : T b → f x ≡ x) →
+      cong f (refl x) ≡
+      trans (f≡id px) (trans (refl x) $ sym (f≡id px′))
+    helper x false px _ f≡id = ⊥-elim px
+    helper x true  _  _ f≡id =
+      cong f (refl x)                                 ≡⟨ cong-refl f ⟩
+      refl (f x)                                      ≡⟨ sym $ trans-symʳ _ ⟩
+      trans (f≡id _) (sym (f≡id _))                   ≡⟨ cong (trans (f≡id _)) $ sym $ trans-reflˡ _ ⟩∎
+      trans (f≡id _) (trans (refl x) $ sym (f≡id _))  ∎
+
+  -- The following lemma is Proposition 2 from "Generalizations of
+  -- Hedberg's Theorem" by Kraus, Escardó, Coquand and Altenkirch.
+
+  subst-in-terms-of-trans-and-cong :
+    ∀ {a b} {A : Set a} {B : Set b} {f g : A → B} {x y}
+      {x≡y : x ≡ y} {fx≡gx : f x ≡ g x} →
+    subst (λ z → f z ≡ g z) x≡y fx≡gx ≡
+    trans (sym (cong f x≡y)) (trans fx≡gx (cong g x≡y))
+  subst-in-terms-of-trans-and-cong {f = f} {g} = elim
+    (λ {x y} x≡y →
+         (fx≡gx : f x ≡ g x) →
+         subst (λ z → f z ≡ g z) x≡y fx≡gx ≡
+         trans (sym (cong f x≡y)) (trans fx≡gx (cong g x≡y)))
+    (λ x fx≡gx →
+         subst (λ z → f z ≡ g z) (refl x) fx≡gx                         ≡⟨ subst-refl _ _ ⟩
+         fx≡gx                                                          ≡⟨ sym $ trans-reflˡ _ ⟩
+         trans (refl (f x)) fx≡gx                                       ≡⟨ sym $ cong₂ trans sym-refl (trans-reflʳ _)  ⟩
+         trans (sym (refl (f x))) (trans fx≡gx (refl (g x)))            ≡⟨ sym $ cong₂ (λ p q → trans (sym p) (trans _ q))
+                                                                                       (cong-refl _) (cong-refl _) ⟩∎
+         trans (sym (cong f (refl x))) (trans fx≡gx (cong g (refl x)))  ∎ )
+    _
+    _
+
+  -- The following lemma is basically Theorem 2.11.5 from the HoTT
+  -- book (the book's lemma gives an equivalence between equality
+  -- types, rather than an equality between equality types).
+
+  [subst≡]≡[trans≡trans] :
+    ∀ {a} {A : Set a} {x y : A} {p : x ≡ y} {q : x ≡ x} {r : y ≡ y} →
+    (subst (λ z → z ≡ z) p q ≡ r)
+      ≡
+    (trans q p ≡ trans p r)
+  [subst≡]≡[trans≡trans] {p = p} {q} {r} = elim
+    (λ {x y} p → {q : x ≡ x} {r : y ≡ y} →
+                 (subst (λ z → z ≡ z) p q ≡ r)
+                   ≡
+                 (trans q p ≡ trans p r))
+    (λ x {q r} →
+       subst (λ z → z ≡ z) (refl x) q ≡ r   ≡⟨ cong (_≡ _) (subst-refl (λ z → z ≡ z) _) ⟩
+       q ≡ r                                ≡⟨ sym $ cong₂ _≡_ (trans-reflʳ _) (trans-reflˡ _) ⟩∎
+       trans q (refl x) ≡ trans (refl x) r  ∎)
+    p
hunk ./Equality/Decidable-UIP.agda 31
-abstract
-
-  -- A set with a constant endofunction with a left inverse is proof
-  -- irrelevant.
-
-  irrelevant : ∀ {a} {A : Set a} →
-               (f : ∃ λ (f : A → A) → Constant f) →
-               (∃ λ g → g Left-inverse-of (proj₁ f)) →
-               Proof-irrelevant A
-  irrelevant (f , constant) (g , left-inverse) x y =
-    x        ≡⟨ sym (left-inverse x) ⟩
-    g (f x)  ≡⟨ cong g (constant x y) ⟩
-    g (f y)  ≡⟨ left-inverse y ⟩∎
-    y        ∎
-
-  -- Endofunction families on _≡_ always have left inverses.
-
-  left-inverse :
-    ∀ {a} {A : Set a} (f : (x y : A) → x ≡ y → x ≡ y) →
-    ∀ {x y} → ∃ λ g → g Left-inverse-of f x y
-  left-inverse {A = A} f {x} {y} =
-    (λ x≡y →
-       x  ≡⟨ x≡y ⟩
-       y  ≡⟨ sym (f y y (refl y)) ⟩∎
-       y  ∎) ,
-    elim (λ {x y} x≡y → trans (f x y x≡y) (sym (f y y (refl y))) ≡ x≡y)
-         (λ _ → trans-symʳ _)
-
-  -- A set A has unique identity proofs if there is a family of
-  -- constant endofunctions on _≡_ {A = A}.
-
-  constant⇒UIP :
-    ∀ {a} {A : Set a} →
-    ((x y : A) → ∃ λ (f : x ≡ y → x ≡ y) → Constant f) →
-    Uniqueness-of-identity-proofs A
-  constant⇒UIP constant {x} {y} =
-    irrelevant (constant x y)
-               (left-inverse (λ x y → proj₁ $ constant x y))
-
-  -- Sets which are decidable come with constant endofunctions.
-
-  decidable⇒constant : ∀ {a} {A : Set a} → Dec A →
-                       ∃ λ (f : A → A) → Constant f
-  decidable⇒constant (yes x) = (const x , λ _ _ → refl x)
-  decidable⇒constant (no ¬x) = (id      , λ _ → ⊥-elim ∘ ¬x)
-
-  -- Sets with decidable equality have unique identity proofs.
-
-  decidable⇒UIP : ∀ {a} {A : Set a} →
-    Decidable-equality A → Uniqueness-of-identity-proofs A
-  decidable⇒UIP dec =
-    constant⇒UIP (λ x y → decidable⇒constant (dec x y))
-
-  -- Types with decidable equality are sets.
-
-  decidable⇒set : ∀ {a} {A : Set a} → Decidable-equality A → Is-set A
-  decidable⇒set {A = A} dec =
-    _⇔_.from {To = Uniqueness-of-identity-proofs A}
-             set⇔UIP (decidable⇒UIP dec)
-
-  -- Non-dependent functions with propositional domains are constant.
-
-  propositional-domain⇒constant :
-    ∀ {a b} {A : Set a} {B : Set b} →
-    Is-proposition A → (f : A → B) → Constant f
-  propositional-domain⇒constant A-prop f = λ x y →
-    cong f (_⇔_.to propositional⇔irrelevant A-prop x y)
-
-  -- If there is a propositional, reflexive relation on A, and related
-  -- elements are equal, then A is a set.
-  --
-  -- (The statement of this lemma is one part of the statement of
-  -- Theorem 7.2.2 in "Homotopy Type Theory: Univalent Foundations of
-  -- Mathematics" (first edition).)
-
-  propositional-identity⇒set :
-    ∀ {a b} {A : Set a}
-    (B : A → A → Set b) →
-    (∀ x y → Is-proposition (B x y)) →
-    (∀ x → B x x) →
-    (∀ x y → B x y → x ≡ y) →
-    Is-set A
-  propositional-identity⇒set B B-prop B-refl f =
-    _⇔_.from set⇔UIP $ constant⇒UIP λ x y →
-      (λ eq → f x y (subst (B x) eq (B-refl x))) ,
-      (λ _ _ → propositional-domain⇒constant (B-prop x y) (f x y) _ _)
-
-  -- The following two results come from "Generalizations of Hedberg's
-  -- Theorem" by Kraus, Escardó, Coquand and Altenkirch.
-
-  -- Proposition 3.
-
-  cong-constant :
-    ∀ {a b} {A : Set a} {B : Set b} {f : A → B} {x} {x≡x : x ≡ x} →
-    (c : Constant f) →
-    cong f x≡x ≡ refl (f x)
-  cong-constant {f = f} {x} {x≡x} c =
-    cong f x≡x                   ≡⟨ elim (λ {x y} x≡y →
-                                              cong f x≡y ≡ trans (sym (c x x)) (c x y))
-                                         (λ x →
-            cong f (refl x)                  ≡⟨ cong-refl _ ⟩
-            refl (f x)                       ≡⟨ sym $ trans-symˡ _ ⟩∎
-            trans (sym (c x x)) (c x x)      ∎)
-                                         _ ⟩
-    trans (sym (c x x)) (c x x)  ≡⟨ trans-symˡ _ ⟩∎
-    refl (f x)                   ∎
-
-  -- The "Fixed Point Lemma".
-
-  fixpoint-lemma :
-    ∀ {a} {A : Set a} →
-    (f : A → A) →
-    Constant f →
-    Is-proposition (∃ λ x → f x ≡ x)
-  fixpoint-lemma f constant =
-    _⇔_.from propositional⇔irrelevant λ { (x , fx≡x) (y , fy≡y) →
-      let x≡y = x    ≡⟨ sym fx≡x ⟩
-                f x  ≡⟨ constant x y ⟩
-                f y  ≡⟨ fy≡y ⟩∎
-                y    ∎
-
-          x≡x = x    ≡⟨ sym fx≡x ⟩
-                f x  ≡⟨ subst (λ z → f z ≡ z) (sym x≡y) fy≡y ⟩∎
-                x    ∎
-
-          lemma =
-            subst (λ z → f z ≡ z) x≡x fx≡x                       ≡⟨ subst-in-terms-of-trans-and-cong ⟩
-
-            trans (sym (cong f x≡x)) (trans fx≡x (cong id x≡x))  ≡⟨ cong₂ (λ p q → trans (sym p) (trans _ q))
-                                                                          (cong-constant constant) (sym $ cong-id _) ⟩
-            trans (sym (refl (f x))) (trans fx≡x x≡x)            ≡⟨ cong (λ p → trans p (trans fx≡x x≡x)) sym-refl ⟩
-
-            trans (refl (f x)) (trans fx≡x x≡x)                  ≡⟨ trans-reflˡ _ ⟩
-
-            trans fx≡x x≡x                                       ≡⟨ sym $ trans-assoc _ _ _ ⟩
-
-            trans (trans fx≡x (sym fx≡x))
-                  (subst (λ z → f z ≡ z) (sym x≡y) fy≡y)         ≡⟨ cong (λ p → trans p (subst (λ z → f z ≡ z) (sym x≡y) fy≡y)) $
-                                                                      trans-symʳ _ ⟩
-            trans (refl (f x))
-                  (subst (λ z → f z ≡ z) (sym x≡y) fy≡y)         ≡⟨ trans-reflˡ _ ⟩∎
-
-            subst (λ z → f z ≡ z) (sym x≡y) fy≡y                 ∎
-      in
-      x , fx≡x                                  ≡⟨ Σ-≡,≡→≡ x≡x lemma ⟩
-      x , subst (λ z → f z ≡ z) (sym x≡y) fy≡y  ≡⟨ sym $ Σ-≡,≡→≡ (sym x≡y) (refl _) ⟩∎
-      y , fy≡y                                  ∎ }
+-- A set with a constant endofunction with a left inverse is proof
+-- irrelevant.
+
+irrelevant : ∀ {a} {A : Set a} →
+             (f : ∃ λ (f : A → A) → Constant f) →
+             (∃ λ g → g Left-inverse-of (proj₁ f)) →
+             Proof-irrelevant A
+irrelevant (f , constant) (g , left-inverse) x y =
+  x        ≡⟨ sym (left-inverse x) ⟩
+  g (f x)  ≡⟨ cong g (constant x y) ⟩
+  g (f y)  ≡⟨ left-inverse y ⟩∎
+  y        ∎
+
+-- Endofunction families on _≡_ always have left inverses.
+
+left-inverse :
+  ∀ {a} {A : Set a} (f : (x y : A) → x ≡ y → x ≡ y) →
+  ∀ {x y} → ∃ λ g → g Left-inverse-of f x y
+left-inverse {A = A} f {x} {y} =
+  (λ x≡y →
+     x  ≡⟨ x≡y ⟩
+     y  ≡⟨ sym (f y y (refl y)) ⟩∎
+     y  ∎) ,
+  elim (λ {x y} x≡y → trans (f x y x≡y) (sym (f y y (refl y))) ≡ x≡y)
+       (λ _ → trans-symʳ _)
+
+-- A set A has unique identity proofs if there is a family of
+-- constant endofunctions on _≡_ {A = A}.
+
+constant⇒UIP :
+  ∀ {a} {A : Set a} →
+  ((x y : A) → ∃ λ (f : x ≡ y → x ≡ y) → Constant f) →
+  Uniqueness-of-identity-proofs A
+constant⇒UIP constant {x} {y} =
+  irrelevant (constant x y)
+             (left-inverse (λ x y → proj₁ $ constant x y))
+
+-- Sets which are decidable come with constant endofunctions.
+
+decidable⇒constant : ∀ {a} {A : Set a} → Dec A →
+                     ∃ λ (f : A → A) → Constant f
+decidable⇒constant (yes x) = (const x , λ _ _ → refl x)
+decidable⇒constant (no ¬x) = (id      , λ _ → ⊥-elim ∘ ¬x)
+
+-- Sets with decidable equality have unique identity proofs.
+
+decidable⇒UIP : ∀ {a} {A : Set a} →
+  Decidable-equality A → Uniqueness-of-identity-proofs A
+decidable⇒UIP dec =
+  constant⇒UIP (λ x y → decidable⇒constant (dec x y))
+
+-- Types with decidable equality are sets.
+
+decidable⇒set : ∀ {a} {A : Set a} → Decidable-equality A → Is-set A
+decidable⇒set {A = A} dec =
+  _⇔_.from {To = Uniqueness-of-identity-proofs A}
+           set⇔UIP (decidable⇒UIP dec)
+
+-- Non-dependent functions with propositional domains are constant.
+
+propositional-domain⇒constant :
+  ∀ {a b} {A : Set a} {B : Set b} →
+  Is-proposition A → (f : A → B) → Constant f
+propositional-domain⇒constant A-prop f = λ x y →
+  cong f (_⇔_.to propositional⇔irrelevant A-prop x y)
+
+-- If there is a propositional, reflexive relation on A, and related
+-- elements are equal, then A is a set.
+--
+-- (The statement of this lemma is one part of the statement of
+-- Theorem 7.2.2 in "Homotopy Type Theory: Univalent Foundations of
+-- Mathematics" (first edition).)
+
+propositional-identity⇒set :
+  ∀ {a b} {A : Set a}
+  (B : A → A → Set b) →
+  (∀ x y → Is-proposition (B x y)) →
+  (∀ x → B x x) →
+  (∀ x y → B x y → x ≡ y) →
+  Is-set A
+propositional-identity⇒set B B-prop B-refl f =
+  _⇔_.from set⇔UIP $ constant⇒UIP λ x y →
+    (λ eq → f x y (subst (B x) eq (B-refl x))) ,
+    (λ _ _ → propositional-domain⇒constant (B-prop x y) (f x y) _ _)
+
+-- The following two results come from "Generalizations of Hedberg's
+-- Theorem" by Kraus, Escardó, Coquand and Altenkirch.
+
+-- Proposition 3.
+
+cong-constant :
+  ∀ {a b} {A : Set a} {B : Set b} {f : A → B} {x} {x≡x : x ≡ x} →
+  (c : Constant f) →
+  cong f x≡x ≡ refl (f x)
+cong-constant {f = f} {x} {x≡x} c =
+  cong f x≡x                   ≡⟨ elim (λ {x y} x≡y →
+                                            cong f x≡y ≡ trans (sym (c x x)) (c x y))
+                                       (λ x →
+          cong f (refl x)                  ≡⟨ cong-refl _ ⟩
+          refl (f x)                       ≡⟨ sym $ trans-symˡ _ ⟩∎
+          trans (sym (c x x)) (c x x)      ∎)
+                                       _ ⟩
+  trans (sym (c x x)) (c x x)  ≡⟨ trans-symˡ _ ⟩∎
+  refl (f x)                   ∎
+
+-- The "Fixed Point Lemma".
+
+fixpoint-lemma :
+  ∀ {a} {A : Set a} →
+  (f : A → A) →
+  Constant f →
+  Is-proposition (∃ λ x → f x ≡ x)
+fixpoint-lemma f constant =
+  _⇔_.from propositional⇔irrelevant λ { (x , fx≡x) (y , fy≡y) →
+    let x≡y = x    ≡⟨ sym fx≡x ⟩
+              f x  ≡⟨ constant x y ⟩
+              f y  ≡⟨ fy≡y ⟩∎
+              y    ∎
+
+        x≡x = x    ≡⟨ sym fx≡x ⟩
+              f x  ≡⟨ subst (λ z → f z ≡ z) (sym x≡y) fy≡y ⟩∎
+              x    ∎
+
+        lemma =
+          subst (λ z → f z ≡ z) x≡x fx≡x                       ≡⟨ subst-in-terms-of-trans-and-cong ⟩
+
+          trans (sym (cong f x≡x)) (trans fx≡x (cong id x≡x))  ≡⟨ cong₂ (λ p q → trans (sym p) (trans _ q))
+                                                                        (cong-constant constant) (sym $ cong-id _) ⟩
+          trans (sym (refl (f x))) (trans fx≡x x≡x)            ≡⟨ cong (λ p → trans p (trans fx≡x x≡x)) sym-refl ⟩
+
+          trans (refl (f x)) (trans fx≡x x≡x)                  ≡⟨ trans-reflˡ _ ⟩
+
+          trans fx≡x x≡x                                       ≡⟨ sym $ trans-assoc _ _ _ ⟩
+
+          trans (trans fx≡x (sym fx≡x))
+                (subst (λ z → f z ≡ z) (sym x≡y) fy≡y)         ≡⟨ cong (λ p → trans p (subst (λ z → f z ≡ z) (sym x≡y) fy≡y)) $
+                                                                    trans-symʳ _ ⟩
+          trans (refl (f x))
+                (subst (λ z → f z ≡ z) (sym x≡y) fy≡y)         ≡⟨ trans-reflˡ _ ⟩∎
+
+          subst (λ z → f z ≡ z) (sym x≡y) fy≡y                 ∎
+    in
+    x , fx≡x                                  ≡⟨ Σ-≡,≡→≡ x≡x lemma ⟩
+    x , subst (λ z → f z ≡ z) (sym x≡y) fy≡y  ≡⟨ sym $ Σ-≡,≡→≡ (sym x≡y) (refl _) ⟩∎
+    y , fy≡y                                  ∎ }
hunk ./Equality/Decision-procedures.agda 46
-  abstract
-
-    true≢false : true ≢ false
-    true≢false true≡false = subst T true≡false _
+  true≢false : true ≢ false
+  true≢false true≡false = subst T true≡false _
hunk ./Equality/Decision-procedures.agda 74
-  abstract
-
-    -- Zero is not equal to the successor of any number.
-
-    0≢+ : {n : ℕ} → zero ≢ suc n
-    0≢+ 0≡+ = subst Zero 0≡+ tt
+  -- Zero is not equal to the successor of any number.
+
+  0≢+ : {n : ℕ} → zero ≢ suc n
+  0≢+ 0≡+ = subst Zero 0≡+ tt
hunk ./Equality/Decision-procedures.agda 130
-  abstract
-
-    -- The values inj₁ x and inj₂ y are never equal.
-
-    inj₁≢inj₂ : {x : A} {y : B} → inj₁ x ≢ inj₂ y
-    inj₁≢inj₂ = Bool.true≢false ∘
-                cong {A = A ⊎ B} {B = Bool} [ const true , const false ]
+  -- The values inj₁ x and inj₂ y are never equal.
+
+  inj₁≢inj₂ : {x : A} {y : B} → inj₁ x ≢ inj₂ y
+  inj₁≢inj₂ = Bool.true≢false ∘
+              cong {A = A ⊎ B} {B = Bool} [ const true , const false ]
hunk ./Equality/Decision-procedures.agda 160
-  abstract
-
-    -- The values [] and x ∷ xs are never equal.
-
-    []≢∷ : ∀ {x : A} {xs} → [] ≢ x ∷ xs
-    []≢∷ = Bool.true≢false ∘
-             cong (λ { [] → true; (_ ∷ _) → false })
+  -- The values [] and x ∷ xs are never equal.
+
+  []≢∷ : ∀ {x : A} {xs} → [] ≢ x ∷ xs
+  []≢∷ = Bool.true≢false ∘
+           cong (λ { [] → true; (_ ∷ _) → false })
hunk ./Equality/Decision-procedures.agda 184
-  abstract
-
-    -- An η-like result for the cancellation lemmas.
-
-    unfold-∷ : ∀ {x y : A} {xs ys} (eq : x ∷ xs ≡ y ∷ ys) →
-               eq ≡ cong₂ _∷_ (cancel-∷-head eq) (cancel-∷-tail eq)
-    unfold-∷ {x} eq =
-      eq                                             ≡⟨ sym $ trans-reflʳ eq ⟩
-      trans eq (refl _)                              ≡⟨ sym $ cong (trans eq) sym-refl ⟩
-      trans eq (sym (refl _))                        ≡⟨ sym $ trans-reflˡ _ ⟩
-      trans (refl _) (trans eq (sym (refl _)))       ≡⟨ sym $ cong-roughly-id (λ xs → head? x xs ∷ tail? xs)
-                                                                              non-empty eq tt tt lemma₁ ⟩
-      cong (λ xs → head? x xs ∷ tail? xs) eq         ≡⟨ lemma₂ _∷_ (head? x) tail? eq ⟩∎
-      cong₂ _∷_ (cong (head? x) eq) (cong tail? eq)  ∎
-      where
-      non-empty : List A → Bool
-      non-empty []      = false
-      non-empty (_ ∷ _) = true
-
-      lemma₁ : (xs : List A) →
-               if non-empty xs then ⊤ else ⊥ →
-               head? x xs ∷ tail? xs ≡ xs
-      lemma₁ []      ()
-      lemma₁ (_ ∷ _) _ = refl _
-
-      lemma₂ : {A B C D : Set a} {x y : A}
-               (f : B → C → D) (g : A → B) (h : A → C) →
-               (eq : x ≡ y) →
-               cong (λ x → f (g x) (h x)) eq ≡
-               cong₂ f (cong g eq) (cong h eq)
-      lemma₂ f g h =
-        elim (λ eq → cong (λ x → f (g x) (h x)) eq ≡
-                     cong₂ f (cong g eq) (cong h eq))
-             (λ x → cong (λ x → f (g x) (h x)) (refl x)          ≡⟨ cong-refl (λ x → f (g x) (h x)) ⟩
-                    refl (f (g x) (h x))                         ≡⟨ sym $ cong₂-refl f ⟩
-                    cong₂ f (refl (g x)) (refl (h x))            ≡⟨ sym $ cong₂ (cong₂ f) (cong-refl g) (cong-refl h) ⟩∎
-                    cong₂ f (cong g (refl x)) (cong h (refl x))  ∎)
+  -- An η-like result for the cancellation lemmas.
+
+  unfold-∷ : ∀ {x y : A} {xs ys} (eq : x ∷ xs ≡ y ∷ ys) →
+             eq ≡ cong₂ _∷_ (cancel-∷-head eq) (cancel-∷-tail eq)
+  unfold-∷ {x} eq =
+    eq                                             ≡⟨ sym $ trans-reflʳ eq ⟩
+    trans eq (refl _)                              ≡⟨ sym $ cong (trans eq) sym-refl ⟩
+    trans eq (sym (refl _))                        ≡⟨ sym $ trans-reflˡ _ ⟩
+    trans (refl _) (trans eq (sym (refl _)))       ≡⟨ sym $ cong-roughly-id (λ xs → head? x xs ∷ tail? xs)
+                                                                            non-empty eq tt tt lemma₁ ⟩
+    cong (λ xs → head? x xs ∷ tail? xs) eq         ≡⟨ lemma₂ _∷_ (head? x) tail? eq ⟩∎
+    cong₂ _∷_ (cong (head? x) eq) (cong tail? eq)  ∎
+    where
+    non-empty : List A → Bool
+    non-empty []      = false
+    non-empty (_ ∷ _) = true
+
+    lemma₁ : (xs : List A) →
+             if non-empty xs then ⊤ else ⊥ →
+             head? x xs ∷ tail? xs ≡ xs
+    lemma₁ []      ()
+    lemma₁ (_ ∷ _) _ = refl _
+
+    lemma₂ : {A B C D : Set a} {x y : A}
+             (f : B → C → D) (g : A → B) (h : A → C) →
+             (eq : x ≡ y) →
+             cong (λ x → f (g x) (h x)) eq ≡
+             cong₂ f (cong g eq) (cong h eq)
+    lemma₂ f g h =
+      elim (λ eq → cong (λ x → f (g x) (h x)) eq ≡
+                   cong₂ f (cong g eq) (cong h eq))
+           (λ x → cong (λ x → f (g x) (h x)) (refl x)          ≡⟨ cong-refl (λ x → f (g x) (h x)) ⟩
+                  refl (f (g x) (h x))                         ≡⟨ sym $ cong₂-refl f ⟩
+                  cong₂ f (refl (g x)) (refl (h x))            ≡⟨ sym $ cong₂ (cong₂ f) (cong-refl g) (cong-refl h) ⟩∎
+                  cong₂ f (cong g (refl x)) (cong h (refl x))  ∎)
hunk ./Equality/Groupoid.agda 51
-  abstract
-
-    commutative : (p q : e ≡ e) → p ∘ q ≡ q ∘ p
-    commutative p q =
-      p ∘ q                                 ≡⟨ cong (_∘_ p) (lem₁ _) ⟩
-      p ∘ (ri ∘ li ⁻¹ ∘ q′ ∘ li ∘ ri ⁻¹)    ≡⟨ prove (Trans (Trans (Trans (Trans (Trans (Sym (Lift ri)) (Lift li)) (Lift q′))
-                                                                          (Sym (Lift li))) (Lift ri)) (Lift p))
-                                                     (Trans (Trans (Sym (Lift ri))
-                                                                              (Trans (Trans (Lift li) (Lift q′)) (Sym (Lift li))))
-                                                            (Trans (Lift ri) (Lift p)))
-                                                     (refl _) ⟩
-      (p ∘ ri) ∘ (li ⁻¹ ∘ q′ ∘ li) ∘ ri ⁻¹  ≡⟨ cong₂ (λ p q → p ∘ q ∘ ri ⁻¹) (lem₂ _) (lem₃ _) ⟩
-      (ri ∘ lc p) ∘ rc q′ ∘ ri ⁻¹           ≡⟨ prove (Trans (Trans (Sym (Lift ri)) (Lift (rc q′))) (Trans (Lift (lc p)) (Lift ri)))
-                                                     (Trans (Trans (Sym (Lift ri)) (Trans (Lift (rc q′)) (Lift (lc p)))) (Lift ri))
-                                                     (refl _) ⟩
-      ri ∘ (lc p ∘ rc q′) ∘ ri ⁻¹           ≡⟨ cong (λ p → ri ∘ p ∘ ri ⁻¹) (lem₄ _ _) ⟩
-      ri ∘ (rc q′ ∘ lc p) ∘ ri ⁻¹           ≡⟨ prove (Trans (Trans (Sym (Lift ri)) (Trans (Lift (lc p)) (Lift (rc q′)))) (Lift ri))
-                                                     (Trans (Trans (Trans (Sym (Lift ri)) (Lift (lc p))) (Lift (rc q′))) (Lift ri))
-                                                     (refl _) ⟩
-      ri ∘ rc q′ ∘ (lc p ∘ ri ⁻¹)           ≡⟨ cong₂ (λ p q → ri ∘ p ∘ q) (sym (lem₃ _)) (lem₅ _) ⟩
-      ri ∘ (li ⁻¹ ∘ q′ ∘ li) ∘ (ri ⁻¹ ∘ p)  ≡⟨ prove (Trans (Trans (Trans (Lift p) (Sym (Lift ri)))
-                                                                   (Trans (Trans (Lift li) (Lift q′)) (Sym (Lift li))))
-                                                            (Lift ri))
-                                                     (Trans (Lift p) (Trans (Trans (Trans (Trans (Sym (Lift ri)) (Lift li)) (Lift q′))
-                                                                                   (Sym (Lift li)))
-                                                                            (Lift ri)))
-                                                     (refl _) ⟩
-      (ri ∘ li ⁻¹ ∘ q′ ∘ li ∘ ri ⁻¹) ∘ p    ≡⟨ cong (λ q → q ∘ p) (sym (lem₁ _)) ⟩∎
-      q ∘ p                                 ∎
-      where
-
-      -- Abbreviations.
-
-      li : ∀ {x} → (e ∙ x) ≡ x
-      li = left-identity _
-
-      ri : ∀ {x} → (x ∙ e) ≡ x
-      ri = right-identity _
-
-      q′ : e ≡ e
-      q′ = li ∘ ri ⁻¹ ∘ q ∘ ri ∘ li ⁻¹
-
-      lc : ∀ {x y} → x ≡ y → (x ∙ e) ≡ (y ∙ e)
-      lc = cong (λ x → (x ∙ e))
-
-      rc : ∀ {x y} → x ≡ y → (e ∙ x) ≡ (e ∙ y)
-      rc = cong (λ y → (e ∙ y))
-
-      -- Lemmas.
-
-      lem₁ : (p : e ≡ e) →
-             p ≡ ri ∘ li ⁻¹ ∘ (li ∘ ri ⁻¹ ∘ p ∘ ri ∘ li ⁻¹) ∘ li ∘ ri ⁻¹
-      lem₁ p =
-        p                                                          ≡⟨ prove (Lift p) (Trans (Trans Refl (Lift p)) Refl) (refl _) ⟩
-        refl _ ∘ p ∘ refl _                                        ≡⟨ sym (cong₂ (λ q r → q ∘ p ∘ r)
-                                                                                 (right-inverse _) (right-inverse _)) ⟩
-        (ri ∘ ri ⁻¹) ∘ p ∘ (ri ∘ ri ⁻¹)                            ≡⟨ prove (Trans (Trans (Trans (Sym (Lift ri)) (Lift ri)) (Lift p))
-                                                                                   (Trans (Sym (Lift ri)) (Lift ri)))
-                                                                            (Trans (Trans (Trans (Trans (Trans (Trans
-                                                                               (Sym (Lift ri)) Refl) (Lift ri)) (Lift p))
-                                                                               (Sym (Lift ri))) Refl) (Lift ri))
-                                                                            (refl _) ⟩
-        ri ∘ refl _ ∘ ri ⁻¹ ∘ p ∘ ri ∘ refl _ ∘ ri ⁻¹              ≡⟨ sym (cong₂ (λ q r → ri ∘ q ∘ ri ⁻¹ ∘ p ∘ ri ∘ r ∘ ri ⁻¹)
-                                                                                 (left-inverse _) (left-inverse _)) ⟩
-        ri ∘ (li ⁻¹ ∘ li) ∘ ri ⁻¹ ∘ p ∘ ri ∘ (li ⁻¹ ∘ li) ∘ ri ⁻¹  ≡⟨ prove (Trans (Trans (Trans (Trans (Trans (Trans
-                                                                               (Sym (Lift ri)) (Trans (Lift li) (Sym (Lift li))))
-                                                                               (Lift ri)) (Lift p)) (Sym (Lift ri)))
-                                                                               (Trans (Lift li) (Sym (Lift li)))) (Lift ri))
-                                                                            (Trans (Trans (Trans (Trans
-                                                                               (Sym (Lift ri)) (Lift li))
-                                                                               (Trans (Trans (Trans (Trans
-                                                                                  (Sym (Lift li)) (Lift ri)) (Lift p)) (Sym (Lift ri)))
-                                                                                  (Lift li))) (Sym (Lift li))) (Lift ri))
-                                                                            (refl _) ⟩∎
-        ri ∘ li ⁻¹ ∘ (li ∘ ri ⁻¹ ∘ p ∘ ri ∘ li ⁻¹) ∘ li ∘ ri ⁻¹    ∎
-
-      lem₂ : ∀ {x y} (p : x ≡ y) → p ∘ ri ≡ ri ∘ lc p
-      lem₂ = elim (λ p → p ∘ ri ≡ ri ∘ lc p) λ _ →
-        prove (Trans (Lift ri) Refl)
-              (Trans (Cong (λ x → (x ∙ e)) Refl) (Lift ri))
-              (refl _)
-
-      lem₃ : ∀ {x y} (p : x ≡ y) → li ⁻¹ ∘ p ∘ li ≡ rc p
-      lem₃ = elim (λ p → li ⁻¹ ∘ p ∘ li ≡ rc p) λ x →
-        li ⁻¹ ∘ refl x ∘ li  ≡⟨ prove (Trans (Trans (Lift li) Refl) (Sym (Lift li)))
-                                      (Trans (Lift li) (Sym (Lift li)))
-                                      (refl _) ⟩
-        li ⁻¹ ∘ li           ≡⟨ left-inverse _ ⟩
-        refl (e ∙ x)         ≡⟨ prove Refl (Cong (λ y → (e ∙ y)) Refl) (refl _) ⟩∎
-        rc (refl x)          ∎
-
-      lem₄ : (p q : e ≡ e) → lc p ∘ rc q ≡ rc q ∘ lc p
-      lem₄ p q = elim
-        (λ {x y} x≡y → lc x≡y ∘ cong (λ z → (x ∙ z)) q ≡
-                       cong (λ z → (y ∙ z)) q ∘ lc x≡y)
-        (λ x → prove (Trans (Cong (λ z → x ∙ z) (Lift q))
-                            (Cong (λ x → x ∙ e) Refl))
-                     (Trans (Cong (λ x → x ∙ e) Refl)
-                            (Cong (λ z → x ∙ z) (Lift q)))
-                     (refl _))
-        p
-
-      lem₅ : ∀ {x y} (p : x ≡ y) → lc p ∘ ri ⁻¹ ≡ ri ⁻¹ ∘ p
-      lem₅ = elim (λ p → lc p ∘ ri ⁻¹ ≡ ri ⁻¹ ∘ p) λ _ →
-        prove (Trans (Sym (Lift ri)) (Cong (λ x → (x ∙ e)) Refl))
-              (Trans Refl (Sym (Lift ri)))
-              (refl _)
+  commutative : (p q : e ≡ e) → p ∘ q ≡ q ∘ p
+  commutative p q =
+    p ∘ q                                 ≡⟨ cong (_∘_ p) (lem₁ _) ⟩
+    p ∘ (ri ∘ li ⁻¹ ∘ q′ ∘ li ∘ ri ⁻¹)    ≡⟨ prove (Trans (Trans (Trans (Trans (Trans (Sym (Lift ri)) (Lift li)) (Lift q′))
+                                                                        (Sym (Lift li))) (Lift ri)) (Lift p))
+                                                   (Trans (Trans (Sym (Lift ri))
+                                                                            (Trans (Trans (Lift li) (Lift q′)) (Sym (Lift li))))
+                                                          (Trans (Lift ri) (Lift p)))
+                                                   (refl _) ⟩
+    (p ∘ ri) ∘ (li ⁻¹ ∘ q′ ∘ li) ∘ ri ⁻¹  ≡⟨ cong₂ (λ p q → p ∘ q ∘ ri ⁻¹) (lem₂ _) (lem₃ _) ⟩
+    (ri ∘ lc p) ∘ rc q′ ∘ ri ⁻¹           ≡⟨ prove (Trans (Trans (Sym (Lift ri)) (Lift (rc q′))) (Trans (Lift (lc p)) (Lift ri)))
+                                                   (Trans (Trans (Sym (Lift ri)) (Trans (Lift (rc q′)) (Lift (lc p)))) (Lift ri))
+                                                   (refl _) ⟩
+    ri ∘ (lc p ∘ rc q′) ∘ ri ⁻¹           ≡⟨ cong (λ p → ri ∘ p ∘ ri ⁻¹) (lem₄ _ _) ⟩
+    ri ∘ (rc q′ ∘ lc p) ∘ ri ⁻¹           ≡⟨ prove (Trans (Trans (Sym (Lift ri)) (Trans (Lift (lc p)) (Lift (rc q′)))) (Lift ri))
+                                                   (Trans (Trans (Trans (Sym (Lift ri)) (Lift (lc p))) (Lift (rc q′))) (Lift ri))
+                                                   (refl _) ⟩
+    ri ∘ rc q′ ∘ (lc p ∘ ri ⁻¹)           ≡⟨ cong₂ (λ p q → ri ∘ p ∘ q) (sym (lem₃ _)) (lem₅ _) ⟩
+    ri ∘ (li ⁻¹ ∘ q′ ∘ li) ∘ (ri ⁻¹ ∘ p)  ≡⟨ prove (Trans (Trans (Trans (Lift p) (Sym (Lift ri)))
+                                                                 (Trans (Trans (Lift li) (Lift q′)) (Sym (Lift li))))
+                                                          (Lift ri))
+                                                   (Trans (Lift p) (Trans (Trans (Trans (Trans (Sym (Lift ri)) (Lift li)) (Lift q′))
+                                                                                 (Sym (Lift li)))
+                                                                          (Lift ri)))
+                                                   (refl _) ⟩
+    (ri ∘ li ⁻¹ ∘ q′ ∘ li ∘ ri ⁻¹) ∘ p    ≡⟨ cong (λ q → q ∘ p) (sym (lem₁ _)) ⟩∎
+    q ∘ p                                 ∎
+    where
+
+    -- Abbreviations.
+
+    li : ∀ {x} → (e ∙ x) ≡ x
+    li = left-identity _
+
+    ri : ∀ {x} → (x ∙ e) ≡ x
+    ri = right-identity _
+
+    q′ : e ≡ e
+    q′ = li ∘ ri ⁻¹ ∘ q ∘ ri ∘ li ⁻¹
+
+    lc : ∀ {x y} → x ≡ y → (x ∙ e) ≡ (y ∙ e)
+    lc = cong (λ x → (x ∙ e))
+
+    rc : ∀ {x y} → x ≡ y → (e ∙ x) ≡ (e ∙ y)
+    rc = cong (λ y → (e ∙ y))
+
+    -- Lemmas.
+
+    lem₁ : (p : e ≡ e) →
+           p ≡ ri ∘ li ⁻¹ ∘ (li ∘ ri ⁻¹ ∘ p ∘ ri ∘ li ⁻¹) ∘ li ∘ ri ⁻¹
+    lem₁ p =
+      p                                                          ≡⟨ prove (Lift p) (Trans (Trans Refl (Lift p)) Refl) (refl _) ⟩
+      refl _ ∘ p ∘ refl _                                        ≡⟨ sym (cong₂ (λ q r → q ∘ p ∘ r)
+                                                                               (right-inverse _) (right-inverse _)) ⟩
+      (ri ∘ ri ⁻¹) ∘ p ∘ (ri ∘ ri ⁻¹)                            ≡⟨ prove (Trans (Trans (Trans (Sym (Lift ri)) (Lift ri)) (Lift p))
+                                                                                 (Trans (Sym (Lift ri)) (Lift ri)))
+                                                                          (Trans (Trans (Trans (Trans (Trans (Trans
+                                                                             (Sym (Lift ri)) Refl) (Lift ri)) (Lift p))
+                                                                             (Sym (Lift ri))) Refl) (Lift ri))
+                                                                          (refl _) ⟩
+      ri ∘ refl _ ∘ ri ⁻¹ ∘ p ∘ ri ∘ refl _ ∘ ri ⁻¹              ≡⟨ sym (cong₂ (λ q r → ri ∘ q ∘ ri ⁻¹ ∘ p ∘ ri ∘ r ∘ ri ⁻¹)
+                                                                               (left-inverse _) (left-inverse _)) ⟩
+      ri ∘ (li ⁻¹ ∘ li) ∘ ri ⁻¹ ∘ p ∘ ri ∘ (li ⁻¹ ∘ li) ∘ ri ⁻¹  ≡⟨ prove (Trans (Trans (Trans (Trans (Trans (Trans
+                                                                             (Sym (Lift ri)) (Trans (Lift li) (Sym (Lift li))))
+                                                                             (Lift ri)) (Lift p)) (Sym (Lift ri)))
+                                                                             (Trans (Lift li) (Sym (Lift li)))) (Lift ri))
+                                                                          (Trans (Trans (Trans (Trans
+                                                                             (Sym (Lift ri)) (Lift li))
+                                                                             (Trans (Trans (Trans (Trans
+                                                                                (Sym (Lift li)) (Lift ri)) (Lift p)) (Sym (Lift ri)))
+                                                                                (Lift li))) (Sym (Lift li))) (Lift ri))
+                                                                          (refl _) ⟩∎
+      ri ∘ li ⁻¹ ∘ (li ∘ ri ⁻¹ ∘ p ∘ ri ∘ li ⁻¹) ∘ li ∘ ri ⁻¹    ∎
+
+    lem₂ : ∀ {x y} (p : x ≡ y) → p ∘ ri ≡ ri ∘ lc p
+    lem₂ = elim (λ p → p ∘ ri ≡ ri ∘ lc p) λ _ →
+      prove (Trans (Lift ri) Refl)
+            (Trans (Cong (λ x → (x ∙ e)) Refl) (Lift ri))
+            (refl _)
+
+    lem₃ : ∀ {x y} (p : x ≡ y) → li ⁻¹ ∘ p ∘ li ≡ rc p
+    lem₃ = elim (λ p → li ⁻¹ ∘ p ∘ li ≡ rc p) λ x →
+      li ⁻¹ ∘ refl x ∘ li  ≡⟨ prove (Trans (Trans (Lift li) Refl) (Sym (Lift li)))
+                                    (Trans (Lift li) (Sym (Lift li)))
+                                    (refl _) ⟩
+      li ⁻¹ ∘ li           ≡⟨ left-inverse _ ⟩
+      refl (e ∙ x)         ≡⟨ prove Refl (Cong (λ y → (e ∙ y)) Refl) (refl _) ⟩∎
+      rc (refl x)          ∎
+
+    lem₄ : (p q : e ≡ e) → lc p ∘ rc q ≡ rc q ∘ lc p
+    lem₄ p q = elim
+      (λ {x y} x≡y → lc x≡y ∘ cong (λ z → (x ∙ z)) q ≡
+                     cong (λ z → (y ∙ z)) q ∘ lc x≡y)
+      (λ x → prove (Trans (Cong (λ z → x ∙ z) (Lift q))
+                          (Cong (λ x → x ∙ e) Refl))
+                   (Trans (Cong (λ x → x ∙ e) Refl)
+                          (Cong (λ z → x ∙ z) (Lift q)))
+                   (refl _))
+      p
+
+    lem₅ : ∀ {x y} (p : x ≡ y) → lc p ∘ ri ⁻¹ ≡ ri ⁻¹ ∘ p
+    lem₅ = elim (λ p → lc p ∘ ri ⁻¹ ≡ ri ⁻¹ ∘ p) λ _ →
+      prove (Trans (Sym (Lift ri)) (Cong (λ x → (x ∙ e)) Refl))
+            (Trans Refl (Sym (Lift ri)))
+            (refl _)
hunk ./Equality/Tactic.agda 111
-  abstract
-
-    lift-correct : ∀ {a} {A : Set a} {x y : A}
-                   (x≡y : x ≡ y) → x≡y ≡ ⟦ lift x≡y ⟧S
-    lift-correct x≡y =
-      x≡y                           ≡⟨ cong-id _ ⟩
-      cong id x≡y                   ≡⟨ sym (trans-reflʳ _) ⟩∎
-      trans (cong id x≡y) (refl _)  ∎
+  lift-correct : ∀ {a} {A : Set a} {x y : A}
+                 (x≡y : x ≡ y) → x≡y ≡ ⟦ lift x≡y ⟧S
+  lift-correct x≡y =
+    x≡y                           ≡⟨ cong-id _ ⟩
+    cong id x≡y                   ≡⟨ sym (trans-reflʳ _) ⟩∎
+    trans (cong id x≡y) (refl _)  ∎
hunk ./Equality/Tactic.agda 123
-  abstract
-
-    snoc-correct :
-      ∀ {a} {A : Set a} {x y z : A} {x≡y y≡z}
-      (z≈y : EqS upper z y) (y≈x : EqS middle y x) →
-      sym y≡z ≡ ⟦ z≈y ⟧S → sym x≡y ≡ ⟦ y≈x ⟧S →
-      sym (trans x≡y y≡z) ≡ ⟦ snoc z≈y y≈x ⟧S
-    snoc-correct {x≡y = x≡y} {y≡z} Refl y≈z h₁ h₂ =
-      sym (trans x≡y y≡z)        ≡⟨ sym-trans _ _ ⟩
-      trans (sym y≡z) (sym x≡y)  ≡⟨ cong₂ trans h₁ h₂ ⟩
-      trans (refl _) ⟦ y≈z ⟧S    ≡⟨ trans-reflˡ _ ⟩
-      ⟦ y≈z ⟧S                   ≡⟨ sym (trans-reflʳ _) ⟩∎
-      trans ⟦ y≈z ⟧S (refl _)    ∎
-    snoc-correct {x≡y = x≡y} {y≡z} (Cons x≈y y≈z) z≈u h₁ h₂ =
-      sym (trans x≡y y≡z)                                    ≡⟨ sym-trans _ _ ⟩
-      trans (sym y≡z) (sym x≡y)                              ≡⟨ cong₂ trans h₁ (refl (sym x≡y)) ⟩
-      trans (trans ⟦ x≈y ⟧S ⟦ y≈z ⟧S) (sym x≡y)              ≡⟨ trans-assoc _ _ _ ⟩
-      trans ⟦ x≈y ⟧S (trans ⟦ y≈z ⟧S (sym x≡y))              ≡⟨ cong (trans ⟦ x≈y ⟧S) $
-                                                                  cong₂ trans (sym (sym-sym ⟦ y≈z ⟧S)) (refl (sym x≡y)) ⟩
-      trans ⟦ x≈y ⟧S (trans (sym (sym ⟦ y≈z ⟧S)) (sym x≡y))  ≡⟨ cong (trans _) $ sym (sym-trans x≡y (sym ⟦ y≈z ⟧S)) ⟩
-      trans ⟦ x≈y ⟧S (sym (trans x≡y (sym ⟦ y≈z ⟧S)))        ≡⟨ cong (trans _) $ snoc-correct y≈z z≈u (sym-sym _) h₂ ⟩∎
-      trans ⟦ x≈y ⟧S ⟦ snoc y≈z z≈u ⟧S                       ∎
+  snoc-correct :
+    ∀ {a} {A : Set a} {x y z : A} {x≡y y≡z}
+    (z≈y : EqS upper z y) (y≈x : EqS middle y x) →
+    sym y≡z ≡ ⟦ z≈y ⟧S → sym x≡y ≡ ⟦ y≈x ⟧S →
+    sym (trans x≡y y≡z) ≡ ⟦ snoc z≈y y≈x ⟧S
+  snoc-correct {x≡y = x≡y} {y≡z} Refl y≈z h₁ h₂ =
+    sym (trans x≡y y≡z)        ≡⟨ sym-trans _ _ ⟩
+    trans (sym y≡z) (sym x≡y)  ≡⟨ cong₂ trans h₁ h₂ ⟩
+    trans (refl _) ⟦ y≈z ⟧S    ≡⟨ trans-reflˡ _ ⟩
+    ⟦ y≈z ⟧S                   ≡⟨ sym (trans-reflʳ _) ⟩∎
+    trans ⟦ y≈z ⟧S (refl _)    ∎
+  snoc-correct {x≡y = x≡y} {y≡z} (Cons x≈y y≈z) z≈u h₁ h₂ =
+    sym (trans x≡y y≡z)                                    ≡⟨ sym-trans _ _ ⟩
+    trans (sym y≡z) (sym x≡y)                              ≡⟨ cong₂ trans h₁ (refl (sym x≡y)) ⟩
+    trans (trans ⟦ x≈y ⟧S ⟦ y≈z ⟧S) (sym x≡y)              ≡⟨ trans-assoc _ _ _ ⟩
+    trans ⟦ x≈y ⟧S (trans ⟦ y≈z ⟧S (sym x≡y))              ≡⟨ cong (trans ⟦ x≈y ⟧S) $
+                                                                cong₂ trans (sym (sym-sym ⟦ y≈z ⟧S)) (refl (sym x≡y)) ⟩
+    trans ⟦ x≈y ⟧S (trans (sym (sym ⟦ y≈z ⟧S)) (sym x≡y))  ≡⟨ cong (trans _) $ sym (sym-trans x≡y (sym ⟦ y≈z ⟧S)) ⟩
+    trans ⟦ x≈y ⟧S (sym (trans x≡y (sym ⟦ y≈z ⟧S)))        ≡⟨ cong (trans _) $ snoc-correct y≈z z≈u (sym-sym _) h₂ ⟩∎
+    trans ⟦ x≈y ⟧S ⟦ snoc y≈z z≈u ⟧S                       ∎
hunk ./Equality/Tactic.agda 149
-  abstract
-
-    append-correct :
-      ∀ {a} {A : Set a} {x y z : A} {x≡y y≡z}
-      (x≈y : EqS upper x y) (y≈z : EqS upper y z) →
-      x≡y ≡ ⟦ x≈y ⟧S → y≡z ≡ ⟦ y≈z ⟧S →
-      trans x≡y y≡z ≡ ⟦ append x≈y y≈z ⟧S
-    append-correct {x≡y = x≡y} {y≡z} Refl x≈y h₁ h₂ =
-      trans x≡y y≡z            ≡⟨ cong₂ trans h₁ h₂ ⟩
-      trans (refl _) ⟦ x≈y ⟧S  ≡⟨ trans-reflˡ _ ⟩∎
-      ⟦ x≈y ⟧S                 ∎
-    append-correct {x≡y = x≡z} {z≡u} (Cons x≈y y≈z) z≈u h₁ h₂ =
-      trans x≡z z≡u                        ≡⟨ cong₂ trans h₁ (refl z≡u) ⟩
-      trans (trans ⟦ x≈y ⟧S ⟦ y≈z ⟧S) z≡u  ≡⟨ trans-assoc _ _ _ ⟩
-      trans ⟦ x≈y ⟧S (trans ⟦ y≈z ⟧S z≡u)  ≡⟨ cong (trans _) $ append-correct y≈z z≈u (refl _) h₂ ⟩∎
-      trans ⟦ x≈y ⟧S ⟦ append y≈z z≈u ⟧S   ∎
+  append-correct :
+    ∀ {a} {A : Set a} {x y z : A} {x≡y y≡z}
+    (x≈y : EqS upper x y) (y≈z : EqS upper y z) →
+    x≡y ≡ ⟦ x≈y ⟧S → y≡z ≡ ⟦ y≈z ⟧S →
+    trans x≡y y≡z ≡ ⟦ append x≈y y≈z ⟧S
+  append-correct {x≡y = x≡y} {y≡z} Refl x≈y h₁ h₂ =
+    trans x≡y y≡z            ≡⟨ cong₂ trans h₁ h₂ ⟩
+    trans (refl _) ⟦ x≈y ⟧S  ≡⟨ trans-reflˡ _ ⟩∎
+    ⟦ x≈y ⟧S                 ∎
+  append-correct {x≡y = x≡z} {z≡u} (Cons x≈y y≈z) z≈u h₁ h₂ =
+    trans x≡z z≡u                        ≡⟨ cong₂ trans h₁ (refl z≡u) ⟩
+    trans (trans ⟦ x≈y ⟧S ⟦ y≈z ⟧S) z≡u  ≡⟨ trans-assoc _ _ _ ⟩
+    trans ⟦ x≈y ⟧S (trans ⟦ y≈z ⟧S z≡u)  ≡⟨ cong (trans _) $ append-correct y≈z z≈u (refl _) h₂ ⟩∎
+    trans ⟦ x≈y ⟧S ⟦ append y≈z z≈u ⟧S   ∎
hunk ./Equality/Tactic.agda 169
-  abstract
-
-    map-sym-correct :
-      ∀ {a} {A : Set a} {x y : A} {x≡y}
-      (x≈y : EqS middle x y) →
-      x≡y ≡ ⟦ x≈y ⟧S → sym x≡y ≡ ⟦ map-sym x≈y ⟧S
-    map-sym-correct {x≡y = x≡y} (No-Sym x≈y) h =
-      sym x≡y       ≡⟨ cong sym h ⟩∎
-      sym ⟦ x≈y ⟧S  ∎
-    map-sym-correct {x≡y = x≡y} (Sym x≈y) h =
-      sym x≡y             ≡⟨ cong sym h ⟩
-      sym (sym ⟦ x≈y ⟧S)  ≡⟨ sym-sym _ ⟩∎
-      ⟦ x≈y ⟧S            ∎
+  map-sym-correct :
+    ∀ {a} {A : Set a} {x y : A} {x≡y}
+    (x≈y : EqS middle x y) →
+    x≡y ≡ ⟦ x≈y ⟧S → sym x≡y ≡ ⟦ map-sym x≈y ⟧S
+  map-sym-correct {x≡y = x≡y} (No-Sym x≈y) h =
+    sym x≡y       ≡⟨ cong sym h ⟩∎
+    sym ⟦ x≈y ⟧S  ∎
+  map-sym-correct {x≡y = x≡y} (Sym x≈y) h =
+    sym x≡y             ≡⟨ cong sym h ⟩
+    sym (sym ⟦ x≈y ⟧S)  ≡⟨ sym-sym _ ⟩∎
+    ⟦ x≈y ⟧S            ∎
hunk ./Equality/Tactic.agda 186
-  abstract
-
-    reverse-correct :
-      ∀ {a} {A : Set a} {x y : A} {x≡y}
-      (x≈y : EqS upper x y) →
-      x≡y ≡ ⟦ x≈y ⟧S → sym x≡y ≡ ⟦ reverse x≈y ⟧S
-    reverse-correct {x≡y = x≡y} Refl h =
-      sym x≡y       ≡⟨ cong sym h ⟩
-      sym (refl _)  ≡⟨ sym-refl ⟩∎
-      refl _        ∎
-    reverse-correct {x≡y = x≡y} (Cons x≈y y≈z) h =
-      sym x≡y                                ≡⟨ cong sym h ⟩
-      sym (trans ⟦ x≈y ⟧S ⟦ y≈z ⟧S)          ≡⟨ snoc-correct (reverse y≈z) _
-                                                             (reverse-correct y≈z (refl _))
-                                                             (map-sym-correct x≈y (refl _)) ⟩∎
-      ⟦ snoc (reverse y≈z) (map-sym x≈y) ⟧S  ∎
+  reverse-correct :
+    ∀ {a} {A : Set a} {x y : A} {x≡y}
+    (x≈y : EqS upper x y) →
+    x≡y ≡ ⟦ x≈y ⟧S → sym x≡y ≡ ⟦ reverse x≈y ⟧S
+  reverse-correct {x≡y = x≡y} Refl h =
+    sym x≡y       ≡⟨ cong sym h ⟩
+    sym (refl _)  ≡⟨ sym-refl ⟩∎
+    refl _        ∎
+  reverse-correct {x≡y = x≡y} (Cons x≈y y≈z) h =
+    sym x≡y                                ≡⟨ cong sym h ⟩
+    sym (trans ⟦ x≈y ⟧S ⟦ y≈z ⟧S)          ≡⟨ snoc-correct (reverse y≈z) _
+                                                           (reverse-correct y≈z (refl _))
+                                                           (map-sym-correct x≈y (refl _)) ⟩∎
+    ⟦ snoc (reverse y≈z) (map-sym x≈y) ⟧S  ∎
hunk ./Equality/Tactic.agda 211
-  abstract
-
-    map-cong-correct :
-      ∀ {ℓ a} {A B : Set a} {x y : A} (f : A → B) {x≡y}
-      (x≈y : EqS ℓ x y) →
-      x≡y ≡ ⟦ x≈y ⟧S → cong f x≡y ≡ ⟦ map-cong f x≈y ⟧S
-    map-cong-correct {lower}  f {gx≡gy} (Cong g x≡y)   h = cong f gx≡gy         ≡⟨ cong (cong f) h ⟩
-                                                           cong f (cong g x≡y)  ≡⟨ cong-∘ f g _ ⟩∎
-                                                           cong (f ∘ g) x≡y     ∎
-    map-cong-correct {middle} f {x≡y}   (No-Sym x≈y)   h = cong f x≡y           ≡⟨ map-cong-correct f x≈y h ⟩∎
-                                                           ⟦ map-cong f x≈y ⟧S  ∎
-    map-cong-correct {middle} f {x≡y}   (Sym    y≈x)   h = cong f x≡y               ≡⟨ cong (cong f) h ⟩
-                                                           cong f (sym ⟦ y≈x ⟧S)    ≡⟨ cong-sym f _ ⟩
-                                                           sym (cong f ⟦ y≈x ⟧S)    ≡⟨ cong sym (map-cong-correct f y≈x (refl _)) ⟩∎
-                                                           sym ⟦ map-cong f y≈x ⟧S  ∎
-    map-cong-correct {upper}  f {x≡y}    Refl          h = cong f x≡y       ≡⟨ cong (cong f) h ⟩
-                                                           cong f (refl _)  ≡⟨ cong-refl f ⟩∎
-                                                           refl _           ∎
-    map-cong-correct {upper}  f {x≡y}   (Cons x≈y y≈z) h =
-      cong f x≡y                                     ≡⟨ cong (cong f) h ⟩
-      cong f (trans ⟦ x≈y ⟧S ⟦ y≈z ⟧S)               ≡⟨ cong-trans f _ _ ⟩
-      trans (cong f ⟦ x≈y ⟧S) (cong f ⟦ y≈z ⟧S)      ≡⟨ cong₂ trans (map-cong-correct f x≈y (refl _))
-                                                                    (map-cong-correct f y≈z (refl _)) ⟩∎
-      trans ⟦ map-cong f x≈y ⟧S ⟦ map-cong f y≈z ⟧S  ∎
+  map-cong-correct :
+    ∀ {ℓ a} {A B : Set a} {x y : A} (f : A → B) {x≡y}
+    (x≈y : EqS ℓ x y) →
+    x≡y ≡ ⟦ x≈y ⟧S → cong f x≡y ≡ ⟦ map-cong f x≈y ⟧S
+  map-cong-correct {lower}  f {gx≡gy} (Cong g x≡y)   h = cong f gx≡gy         ≡⟨ cong (cong f) h ⟩
+                                                         cong f (cong g x≡y)  ≡⟨ cong-∘ f g _ ⟩∎
+                                                         cong (f ∘ g) x≡y     ∎
+  map-cong-correct {middle} f {x≡y}   (No-Sym x≈y)   h = cong f x≡y           ≡⟨ map-cong-correct f x≈y h ⟩∎
+                                                         ⟦ map-cong f x≈y ⟧S  ∎
+  map-cong-correct {middle} f {x≡y}   (Sym    y≈x)   h = cong f x≡y               ≡⟨ cong (cong f) h ⟩
+                                                         cong f (sym ⟦ y≈x ⟧S)    ≡⟨ cong-sym f _ ⟩
+                                                         sym (cong f ⟦ y≈x ⟧S)    ≡⟨ cong sym (map-cong-correct f y≈x (refl _)) ⟩∎
+                                                         sym ⟦ map-cong f y≈x ⟧S  ∎
+  map-cong-correct {upper}  f {x≡y}    Refl          h = cong f x≡y       ≡⟨ cong (cong f) h ⟩
+                                                         cong f (refl _)  ≡⟨ cong-refl f ⟩∎
+                                                         refl _           ∎
+  map-cong-correct {upper}  f {x≡y}   (Cons x≈y y≈z) h =
+    cong f x≡y                                     ≡⟨ cong (cong f) h ⟩
+    cong f (trans ⟦ x≈y ⟧S ⟦ y≈z ⟧S)               ≡⟨ cong-trans f _ _ ⟩
+    trans (cong f ⟦ x≈y ⟧S) (cong f ⟦ y≈z ⟧S)      ≡⟨ cong₂ trans (map-cong-correct f x≈y (refl _))
+                                                                  (map-cong-correct f y≈z (refl _)) ⟩∎
+    trans ⟦ map-cong f x≈y ⟧S ⟦ map-cong f y≈z ⟧S  ∎
hunk ./Equality/Tactic.agda 244
-abstract
-
-  simplify-correct :
-    ∀ {a} {A : Set a} {x y : A}
-    (x≈y : Eq x y) → ⟦ x≈y ⟧ ≡ ⟦ simplify x≈y ⟧S
-  simplify-correct (Lift x≡y)      = lift-correct x≡y
-  simplify-correct Refl            = refl _
-  simplify-correct (Sym x≈y)       = reverse-correct (simplify x≈y)
-                                       (simplify-correct x≈y)
-  simplify-correct (Trans x≈y y≈z) = append-correct (simplify x≈y) _
-                                       (simplify-correct x≈y)
-                                       (simplify-correct y≈z)
-  simplify-correct (Cong f x≈y)    = map-cong-correct f (simplify x≈y)
-                                       (simplify-correct x≈y)
+simplify-correct :
+  ∀ {a} {A : Set a} {x y : A}
+  (x≈y : Eq x y) → ⟦ x≈y ⟧ ≡ ⟦ simplify x≈y ⟧S
+simplify-correct (Lift x≡y)      = lift-correct x≡y
+simplify-correct Refl            = refl _
+simplify-correct (Sym x≈y)       = reverse-correct (simplify x≈y)
+                                     (simplify-correct x≈y)
+simplify-correct (Trans x≈y y≈z) = append-correct (simplify x≈y) _
+                                     (simplify-correct x≈y)
+                                     (simplify-correct y≈z)
+simplify-correct (Cong f x≈y)    = map-cong-correct f (simplify x≈y)
+                                     (simplify-correct x≈y)
hunk ./Equality/Tactic.agda 260
-abstract
-
-  -- Simple tactic for proving equality of equality proofs.
-
-  prove : ∀ {a} {A : Set a} {x y : A} (x≡y x≡y′ : Eq x y) →
-          ⟦ simplify x≡y ⟧S ≡ ⟦ simplify x≡y′ ⟧S →
-          ⟦ x≡y ⟧ ≡ ⟦ x≡y′ ⟧
-  prove x≡y x≡y′ hyp =
-    ⟦ x≡y ⟧             ≡⟨ simplify-correct x≡y ⟩
-    ⟦ simplify x≡y  ⟧S  ≡⟨ hyp ⟩
-    ⟦ simplify x≡y′ ⟧S  ≡⟨ sym (simplify-correct x≡y′) ⟩∎
-    ⟦ x≡y′ ⟧            ∎
+-- Simple tactic for proving equality of equality proofs.
+
+prove : ∀ {a} {A : Set a} {x y : A} (x≡y x≡y′ : Eq x y) →
+        ⟦ simplify x≡y ⟧S ≡ ⟦ simplify x≡y′ ⟧S →
+        ⟦ x≡y ⟧ ≡ ⟦ x≡y′ ⟧
+prove x≡y x≡y′ hyp =
+  ⟦ x≡y ⟧             ≡⟨ simplify-correct x≡y ⟩
+  ⟦ simplify x≡y  ⟧S  ≡⟨ hyp ⟩
+  ⟦ simplify x≡y′ ⟧S  ≡⟨ sym (simplify-correct x≡y′) ⟩∎
+  ⟦ x≡y′ ⟧            ∎
hunk ./Equivalence.agda 37
-abstract
-
-  -- Is-equivalence f is a proposition, assuming extensional equality.
-
-  propositional :
-    ∀ {a b} → Extensionality (a ⊔ b) (a ⊔ b) →
-    {A : Set a} {B : Set b} (f : A → B) →
-    Is-proposition (Is-equivalence f)
-  propositional {a} ext f =
-    Π-closure (lower-extensionality a lzero ext) 1 λ _ →
-      Contractible-propositional ext
-
-  -- If the domain is contractible and the codomain is propositional,
-  -- then Is-equivalence f is contractible.
-
-  sometimes-contractible :
-    ∀ {a b} → Extensionality (a ⊔ b) (a ⊔ b) →
-    {A : Set a} {B : Set b} {f : A → B} →
-    Contractible A → Is-proposition B →
-    Contractible (Is-equivalence f)
-  sometimes-contractible {a} ext A-contr B-prop =
-    Π-closure (lower-extensionality a lzero ext) 0 λ _ →
-      cojoin ext (Σ-closure 0 A-contr (λ _ → B-prop _ _))
-
-  -- Is-equivalence f is not always contractible.
-
-  not-always-contractible₁ :
-    ∀ {a b} →
-    ∃ λ (A : Set a) → ∃ λ (B : Set b) → ∃ λ (f : A → B) →
-      Is-proposition A × Contractible B ×
-      ¬ Contractible (Is-equivalence f)
-  not-always-contractible₁ =
-    ⊥ ,
-    ↑ _ ⊤ ,
-    const (lift tt) ,
-    ⊥-propositional ,
-    ↑-closure 0 ⊤-contractible ,
-    λ c → ⊥-elim (proj₁ (proj₁ (proj₁ c (lift tt))))
-
-  not-always-contractible₂ :
-    ∀ {a b} →
-    ∃ λ (A : Set a) → ∃ λ (B : Set b) → ∃ λ (f : A → B) →
-      Contractible A × Is-set B ×
-      ¬ Contractible (Is-equivalence f)
-  not-always-contractible₂ =
-    ↑ _ ⊤ ,
-    ↑ _ Bool ,
-    const (lift true) ,
-    ↑-closure 0 ⊤-contractible ,
-    ↑-closure 2 Bool-set ,
-    λ c → Bool.true≢false (cong lower
-            (proj₂ (proj₁ (proj₁ c (lift false)))))
+-- Is-equivalence f is a proposition, assuming extensional equality.
+
+propositional :
+  ∀ {a b} → Extensionality (a ⊔ b) (a ⊔ b) →
+  {A : Set a} {B : Set b} (f : A → B) →
+  Is-proposition (Is-equivalence f)
+propositional {a} ext f =
+  Π-closure (lower-extensionality a lzero ext) 1 λ _ →
+    Contractible-propositional ext
+
+-- If the domain is contractible and the codomain is propositional,
+-- then Is-equivalence f is contractible.
+
+sometimes-contractible :
+  ∀ {a b} → Extensionality (a ⊔ b) (a ⊔ b) →
+  {A : Set a} {B : Set b} {f : A → B} →
+  Contractible A → Is-proposition B →
+  Contractible (Is-equivalence f)
+sometimes-contractible {a} ext A-contr B-prop =
+  Π-closure (lower-extensionality a lzero ext) 0 λ _ →
+    cojoin ext (Σ-closure 0 A-contr (λ _ → B-prop _ _))
+
+-- Is-equivalence f is not always contractible.
+
+not-always-contractible₁ :
+  ∀ {a b} →
+  ∃ λ (A : Set a) → ∃ λ (B : Set b) → ∃ λ (f : A → B) →
+    Is-proposition A × Contractible B ×
+    ¬ Contractible (Is-equivalence f)
+not-always-contractible₁ =
+  ⊥ ,
+  ↑ _ ⊤ ,
+  const (lift tt) ,
+  ⊥-propositional ,
+  ↑-closure 0 ⊤-contractible ,
+  λ c → ⊥-elim (proj₁ (proj₁ (proj₁ c (lift tt))))
+
+not-always-contractible₂ :
+  ∀ {a b} →
+  ∃ λ (A : Set a) → ∃ λ (B : Set b) → ∃ λ (f : A → B) →
+    Contractible A × Is-set B ×
+    ¬ Contractible (Is-equivalence f)
+not-always-contractible₂ =
+  ↑ _ ⊤ ,
+  ↑ _ Bool ,
+  const (lift true) ,
+  ↑-closure 0 ⊤-contractible ,
+  ↑-closure 2 Bool-set ,
+  λ c → Bool.true≢false (cong lower
+          (proj₂ (proj₁ (proj₁ c (lift false)))))
hunk ./Equivalence.agda 100
-abstract
-
-  -- If Σ-map id f is an equivalence, then f is also an equivalence.
-
-  drop-Σ-map-id :
-    ∀ {a b} {A : Set a} {B C : A → Set b} (f : ∀ {x} → B x → C x) →
-    Is-equivalence {A = Σ A B} {B = Σ A C} (Σ-map P.id f) →
-    ∀ x → Is-equivalence (f {x = x})
-  drop-Σ-map-id {b = b} {A} {B} {C} f eq x z =
-    H-level.respects-surjection surj 0 (eq (x , z))
-    where
-    map-f : Σ A B → Σ A C
-    map-f = Σ-map P.id f
-
-    to-P : ∀ {x y} {p : ∃ C} → (x , f y) ≡ p → Set b
-    to-P {y = y} {p} _ = ∃ λ y′ → f y′ ≡ proj₂ p
-
-    to : map-f ⁻¹ (x , z) → f ⁻¹ z
-    to ((x′ , y) , eq) = elim¹ to-P (y , refl (f y)) eq
-
-    from : f ⁻¹ z → map-f ⁻¹ (x , z)
-    from (y , eq) = (x , y) , cong (_,_ x) eq
-
-    to∘from : ∀ p → to (from p) ≡ p
-    to∘from (y , eq) = elim¹
-      (λ {z′} (eq : f y ≡ z′) →
-         _≡_ {A = ∃ λ (y : B x) → f y ≡ z′}
-             (elim¹ to-P (y , refl (f y)) (cong (_,_ x) eq))
-             (y , eq))
-      (elim¹ to-P (y , refl (f y)) (cong (_,_ x) (refl (f y)))  ≡⟨ cong (elim¹ to-P (y , refl (f y))) $
-                                                                        cong-refl (_,_ x) {x = f y} ⟩
-       elim¹ to-P (y , refl (f y)) (refl (x , f y))             ≡⟨ elim¹-refl to-P _ ⟩∎
-       (y , refl (f y))                                         ∎)
-      eq
-
-    surj : map-f ⁻¹ (x , z) ↠ f ⁻¹ z
-    surj = record
-      { logical-equivalence = record { to = to; from = from }
-      ; right-inverse-of    = to∘from
-      }
+-- If Σ-map id f is an equivalence, then f is also an equivalence.
+
+drop-Σ-map-id :
+  ∀ {a b} {A : Set a} {B C : A → Set b} (f : ∀ {x} → B x → C x) →
+  Is-equivalence {A = Σ A B} {B = Σ A C} (Σ-map P.id f) →
+  ∀ x → Is-equivalence (f {x = x})
+drop-Σ-map-id {b = b} {A} {B} {C} f eq x z =
+  H-level.respects-surjection surj 0 (eq (x , z))
+  where
+  map-f : Σ A B → Σ A C
+  map-f = Σ-map P.id f
+
+  to-P : ∀ {x y} {p : ∃ C} → (x , f y) ≡ p → Set b
+  to-P {y = y} {p} _ = ∃ λ y′ → f y′ ≡ proj₂ p
+
+  to : map-f ⁻¹ (x , z) → f ⁻¹ z
+  to ((x′ , y) , eq) = elim¹ to-P (y , refl (f y)) eq
+
+  from : f ⁻¹ z → map-f ⁻¹ (x , z)
+  from (y , eq) = (x , y) , cong (_,_ x) eq
+
+  to∘from : ∀ p → to (from p) ≡ p
+  to∘from (y , eq) = elim¹
+    (λ {z′} (eq : f y ≡ z′) →
+       _≡_ {A = ∃ λ (y : B x) → f y ≡ z′}
+           (elim¹ to-P (y , refl (f y)) (cong (_,_ x) eq))
+           (y , eq))
+    (elim¹ to-P (y , refl (f y)) (cong (_,_ x) (refl (f y)))  ≡⟨ cong (elim¹ to-P (y , refl (f y))) $
+                                                                      cong-refl (_,_ x) {x = f y} ⟩
+     elim¹ to-P (y , refl (f y)) (refl (x , f y))             ≡⟨ elim¹-refl to-P _ ⟩∎
+     (y , refl (f y))                                         ∎)
+    eq
+
+  surj : map-f ⁻¹ (x , z) ↠ f ⁻¹ z
+  surj = record
+    { logical-equivalence = record { to = to; from = from }
+    ; right-inverse-of    = to∘from
+    }
hunk ./Equivalence.agda 160
-  abstract
-
-    left-inverse-of : ∀ x → from (to x) ≡ x
-    left-inverse-of x =
-      cong (proj₁ {B = λ x′ → to x′ ≡ to x}) (
-        proj₁ (is-equivalence (to x))  ≡⟨ proj₂ (is-equivalence (to x)) (x , refl (to x)) ⟩∎
-        (x , refl (to x))              ∎)
+  left-inverse-of : ∀ x → from (to x) ≡ x
+  left-inverse-of x =
+    cong (proj₁ {B = λ x′ → to x′ ≡ to x}) (
+      proj₁ (is-equivalence (to x))  ≡⟨ proj₂ (is-equivalence (to x)) (x , refl (to x)) ⟩∎
+      (x , refl (to x))              ∎)
hunk ./Equivalence.agda 181
-  abstract
-
-    -- All preimages of an element under the equivalence are equal.
-
-    irrelevance : ∀ y (p : to ⁻¹ y) → (from y , right-inverse-of y) ≡ p
-    irrelevance y = proj₂ (is-equivalence y)
-
-    -- The two proofs left-inverse-of and right-inverse-of are
-    -- related.
-
-    left-right-lemma :
-      ∀ x → cong to (left-inverse-of x) ≡ right-inverse-of (to x)
-    left-right-lemma x =
-      lemma₁ to _ _ (lemma₂ (irrelevance (to x) (x , refl (to x))))
-      where
-      lemma₁ : {x y : A} (f : A → B) (p : x ≡ y) (q : f x ≡ f y) →
-               refl (f y) ≡ trans (cong f (sym p)) q →
-               cong f p ≡ q
-      lemma₁ f = elim
-        (λ {x y} p → ∀ q → refl (f y) ≡ trans (cong f (sym p)) q →
-                           cong f p ≡ q)
-        (λ x q hyp →
-           cong f (refl x)                  ≡⟨ cong-refl f ⟩
-           refl (f x)                       ≡⟨ hyp ⟩
-           trans (cong f (sym (refl x))) q  ≡⟨ cong (λ p → trans (cong f p) q) sym-refl ⟩
-           trans (cong f (refl x)) q        ≡⟨ cong (λ p → trans p q) (cong-refl f) ⟩
-           trans (refl (f x)) q             ≡⟨ trans-reflˡ _ ⟩∎
-           q                                ∎)
-
-      lemma₂ : ∀ {f : A → B} {y} {f⁻¹y₁ f⁻¹y₂ : f ⁻¹ y}
-               (p : f⁻¹y₁ ≡ f⁻¹y₂) →
-               proj₂ f⁻¹y₂ ≡
-               trans (cong f (sym (cong (proj₁ {B = λ x → f x ≡ y}) p)))
-                     (proj₂ f⁻¹y₁)
-      lemma₂ {f} {y} =
-        let pr = proj₁ {B = λ x → f x ≡ y} in
-        elim {A = f ⁻¹ y}
-          (λ {f⁻¹y₁ f⁻¹y₂} p →
+  -- All preimages of an element under the equivalence are equal.
+
+  irrelevance : ∀ y (p : to ⁻¹ y) → (from y , right-inverse-of y) ≡ p
+  irrelevance y = proj₂ (is-equivalence y)
+
+  -- The two proofs left-inverse-of and right-inverse-of are
+  -- related.
+
+  left-right-lemma :
+    ∀ x → cong to (left-inverse-of x) ≡ right-inverse-of (to x)
+  left-right-lemma x =
+    lemma₁ to _ _ (lemma₂ (irrelevance (to x) (x , refl (to x))))
+    where
+    lemma₁ : {x y : A} (f : A → B) (p : x ≡ y) (q : f x ≡ f y) →
+             refl (f y) ≡ trans (cong f (sym p)) q →
+             cong f p ≡ q
+    lemma₁ f = elim
+      (λ {x y} p → ∀ q → refl (f y) ≡ trans (cong f (sym p)) q →
+                         cong f p ≡ q)
+      (λ x q hyp →
+         cong f (refl x)                  ≡⟨ cong-refl f ⟩
+         refl (f x)                       ≡⟨ hyp ⟩
+         trans (cong f (sym (refl x))) q  ≡⟨ cong (λ p → trans (cong f p) q) sym-refl ⟩
+         trans (cong f (refl x)) q        ≡⟨ cong (λ p → trans p q) (cong-refl f) ⟩
+         trans (refl (f x)) q             ≡⟨ trans-reflˡ _ ⟩∎
+         q                                ∎)
+
+    lemma₂ : ∀ {f : A → B} {y} {f⁻¹y₁ f⁻¹y₂ : f ⁻¹ y}
+             (p : f⁻¹y₁ ≡ f⁻¹y₂) →
hunk ./Equivalence.agda 211
-               trans (cong f (sym (cong pr p))) (proj₂ f⁻¹y₁))
-          (λ f⁻¹y →
-             proj₂ f⁻¹y                                               ≡⟨ sym $ trans-reflˡ _ ⟩
-             trans (refl (f (proj₁ f⁻¹y))) (proj₂ f⁻¹y)               ≡⟨ cong (λ p → trans p (proj₂ f⁻¹y)) (sym (cong-refl f)) ⟩
-             trans (cong f (refl (proj₁ f⁻¹y))) (proj₂ f⁻¹y)          ≡⟨ cong (λ p → trans (cong f p) (proj₂ f⁻¹y)) (sym sym-refl) ⟩
-             trans (cong f (sym (refl (proj₁ f⁻¹y)))) (proj₂ f⁻¹y)    ≡⟨ cong (λ p → trans (cong f (sym p)) (proj₂ f⁻¹y))
-                                                                              (sym (cong-refl pr)) ⟩∎
-             trans (cong f (sym (cong pr (refl f⁻¹y)))) (proj₂ f⁻¹y)  ∎)
-
-    right-left-lemma :
-      ∀ x → cong from (right-inverse-of x) ≡ left-inverse-of (from x)
-    right-left-lemma x = subst
-      (λ x → cong from (right-inverse-of x) ≡ left-inverse-of (from x))
-      (right-inverse-of x)
-      (let y = from x in
-
-       cong from (right-inverse-of (to y))                          ≡⟨ cong (cong from) $ sym $ left-right-lemma y ⟩
-       cong from (cong to (left-inverse-of y))                      ≡⟨ cong-∘ from to _ ⟩
-       cong (from ⊚ to) (left-inverse-of y)                         ≡⟨ cong-roughly-id (from ⊚ to) (λ _ → true) (left-inverse-of y)
-                                                                                       _ _ (λ z _ → left-inverse-of z) ⟩
-       trans (left-inverse-of (from (to y)))
-             (trans (left-inverse-of y) (sym (left-inverse-of y)))  ≡⟨ cong (trans _) $ trans-symʳ _ ⟩
-       trans (left-inverse-of (from (to y))) (refl _)               ≡⟨ trans-reflʳ _ ⟩∎
-       left-inverse-of (from (to y))                                ∎)
+             trans (cong f (sym (cong (proj₁ {B = λ x → f x ≡ y}) p)))
+                   (proj₂ f⁻¹y₁)
+    lemma₂ {f} {y} =
+      let pr = proj₁ {B = λ x → f x ≡ y} in
+      elim {A = f ⁻¹ y}
+        (λ {f⁻¹y₁ f⁻¹y₂} p →
+           proj₂ f⁻¹y₂ ≡
+             trans (cong f (sym (cong pr p))) (proj₂ f⁻¹y₁))
+        (λ f⁻¹y →
+           proj₂ f⁻¹y                                               ≡⟨ sym $ trans-reflˡ _ ⟩
+           trans (refl (f (proj₁ f⁻¹y))) (proj₂ f⁻¹y)               ≡⟨ cong (λ p → trans p (proj₂ f⁻¹y)) (sym (cong-refl f)) ⟩
+           trans (cong f (refl (proj₁ f⁻¹y))) (proj₂ f⁻¹y)          ≡⟨ cong (λ p → trans (cong f p) (proj₂ f⁻¹y)) (sym sym-refl) ⟩
+           trans (cong f (sym (refl (proj₁ f⁻¹y)))) (proj₂ f⁻¹y)    ≡⟨ cong (λ p → trans (cong f (sym p)) (proj₂ f⁻¹y))
+                                                                            (sym (cong-refl pr)) ⟩∎
+           trans (cong f (sym (cong pr (refl f⁻¹y)))) (proj₂ f⁻¹y)  ∎)
+
+  right-left-lemma :
+    ∀ x → cong from (right-inverse-of x) ≡ left-inverse-of (from x)
+  right-left-lemma x = subst
+    (λ x → cong from (right-inverse-of x) ≡ left-inverse-of (from x))
+    (right-inverse-of x)
+    (let y = from x in
+
+     cong from (right-inverse-of (to y))                          ≡⟨ cong (cong from) $ sym $ left-right-lemma y ⟩
+     cong from (cong to (left-inverse-of y))                      ≡⟨ cong-∘ from to _ ⟩
+     cong (from ⊚ to) (left-inverse-of y)                         ≡⟨ cong-roughly-id (from ⊚ to) (λ _ → true) (left-inverse-of y)
+                                                                                     _ _ (λ z _ → left-inverse-of z) ⟩
+     trans (left-inverse-of (from (to y)))
+           (trans (left-inverse-of y) (sym (left-inverse-of y)))  ≡⟨ cong (trans _) $ trans-symʳ _ ⟩
+     trans (left-inverse-of (from (to y))) (refl _)               ≡⟨ trans-reflʳ _ ⟩∎
+     left-inverse-of (from (to y))                                ∎)
hunk ./Equivalence.agda 254
-  abstract
-    is-equivalence : Is-equivalence to
-    is-equivalence = Preimage.bijection⁻¹-contractible A↔B
-
-    right-inverse-of : ∀ x → to (from x) ≡ x
-    right-inverse-of = proj₂ ⊚ proj₁ ⊚ is-equivalence
-
-    irrelevance : ∀ y (p : to ⁻¹ y) → (from y , right-inverse-of y) ≡ p
-    irrelevance = proj₂ ⊚ is-equivalence
+  is-equivalence : Is-equivalence to
+  is-equivalence = Preimage.bijection⁻¹-contractible A↔B
+
+  right-inverse-of : ∀ x → to (from x) ≡ x
+  right-inverse-of = proj₂ ⊚ proj₁ ⊚ is-equivalence
+
+  irrelevance : ∀ y (p : to ⁻¹ y) → (from y , right-inverse-of y) ≡ p
+  irrelevance = proj₂ ⊚ is-equivalence
hunk ./Equivalence.agda 282
-abstract
-
-  subst-is-equivalence :
-    ∀ {a p} {A : Set a} (P : A → Set p) {x y : A} (x≡y : x ≡ y) →
-    Is-equivalence (subst P x≡y)
-  subst-is-equivalence P x≡y =
-    _≃_.is-equivalence (subst-as-equivalence P x≡y)
+subst-is-equivalence :
+  ∀ {a p} {A : Set a} (P : A → Set p) {x y : A} (x≡y : x ≡ y) →
+  Is-equivalence (subst P x≡y)
+subst-is-equivalence P x≡y =
+  _≃_.is-equivalence (subst-as-equivalence P x≡y)
hunk ./Equivalence.agda 301
-  abstract
-
-    irr : ∀ y (p : from ⁻¹ y) → (to y , left-inverse-of y) ≡ p
-    irr y (x , from-x≡y) =
-      Σ-≡,≡→≡ (from-to from-x≡y) (elim¹
-        (λ {y} ≡y → subst (λ z → from z ≡ y)
-                          (trans (cong to (sym ≡y))
-                                 (right-inverse-of x))
-                          (left-inverse-of y) ≡ ≡y)
-        (let lemma =
-               trans (cong to (sym (refl (from x))))
-                     (right-inverse-of x)              ≡⟨ cong (λ eq → trans (cong to eq) (right-inverse-of x)) sym-refl ⟩
-
-               trans (cong to (refl (from x)))
-                     (right-inverse-of x)              ≡⟨ cong (λ eq → trans eq (right-inverse-of x)) $ cong-refl to ⟩
-
-               trans (refl (to (from x)))
-                     (right-inverse-of x)              ≡⟨ trans-reflˡ (right-inverse-of x) ⟩∎
-
-               right-inverse-of x                      ∎
-         in
-
-         subst (λ z → from z ≡ from x)
-               (trans (cong to (sym (refl (from x))))
-                      (right-inverse-of x))
-               (left-inverse-of (from x))                    ≡⟨ cong₂ (subst (λ z → from z ≡ from x))
-                                                                      lemma (sym $ right-left-lemma x) ⟩
-         subst (λ z → from z ≡ from x)
-               (right-inverse-of x)
-               (cong from $ right-inverse-of x)              ≡⟨ subst-∘ (λ z → z ≡ from x) from _ ⟩
-
-         subst (λ z → z ≡ from x)
-               (cong from $ right-inverse-of x)
-               (cong from $ right-inverse-of x)              ≡⟨ cong (λ eq → subst (λ z → z ≡ from x) eq
-                                                                                   (cong from $ right-inverse-of x)) $
-                                                                     sym $ sym-sym _ ⟩
-         subst (λ z → z ≡ from x)
-               (sym $ sym $ cong from $ right-inverse-of x)
-               (cong from $ right-inverse-of x)              ≡⟨ subst-trans _ ⟩
-
-         trans (sym $ cong from $ right-inverse-of x)
-               (cong from $ right-inverse-of x)              ≡⟨ trans-symˡ _ ⟩∎
-
-         refl (from x)                                       ∎)
-        from-x≡y)
+  irr : ∀ y (p : from ⁻¹ y) → (to y , left-inverse-of y) ≡ p
+  irr y (x , from-x≡y) =
+    Σ-≡,≡→≡ (from-to from-x≡y) (elim¹
+      (λ {y} ≡y → subst (λ z → from z ≡ y)
+                        (trans (cong to (sym ≡y))
+                               (right-inverse-of x))
+                        (left-inverse-of y) ≡ ≡y)
+      (let lemma =
+             trans (cong to (sym (refl (from x))))
+                   (right-inverse-of x)              ≡⟨ cong (λ eq → trans (cong to eq) (right-inverse-of x)) sym-refl ⟩
+
+             trans (cong to (refl (from x)))
+                   (right-inverse-of x)              ≡⟨ cong (λ eq → trans eq (right-inverse-of x)) $ cong-refl to ⟩
+
+             trans (refl (to (from x)))
+                   (right-inverse-of x)              ≡⟨ trans-reflˡ (right-inverse-of x) ⟩∎
+
+             right-inverse-of x                      ∎
+       in
+
+       subst (λ z → from z ≡ from x)
+             (trans (cong to (sym (refl (from x))))
+                    (right-inverse-of x))
+             (left-inverse-of (from x))                    ≡⟨ cong₂ (subst (λ z → from z ≡ from x))
+                                                                    lemma (sym $ right-left-lemma x) ⟩
+       subst (λ z → from z ≡ from x)
+             (right-inverse-of x)
+             (cong from $ right-inverse-of x)              ≡⟨ subst-∘ (λ z → z ≡ from x) from _ ⟩
+
+       subst (λ z → z ≡ from x)
+             (cong from $ right-inverse-of x)
+             (cong from $ right-inverse-of x)              ≡⟨ cong (λ eq → subst (λ z → z ≡ from x) eq
+                                                                                 (cong from $ right-inverse-of x)) $
+                                                                   sym $ sym-sym _ ⟩
+       subst (λ z → z ≡ from x)
+             (sym $ sym $ cong from $ right-inverse-of x)
+             (cong from $ right-inverse-of x)              ≡⟨ subst-trans _ ⟩
+
+       trans (sym $ cong from $ right-inverse-of x)
+             (cong from $ right-inverse-of x)              ≡⟨ trans-symˡ _ ⟩∎
+
+       refl (from x)                                       ∎)
+      from-x≡y)
hunk ./Equivalence.agda 359
-  abstract
-    right-inverse-of : ∀ x → to (from x) ≡ x
-    right-inverse-of = _≃_.right-inverse-of f∘g
-
-    irrelevance : ∀ y (p : to ⁻¹ y) → (from y , right-inverse-of y) ≡ p
-    irrelevance = _≃_.irrelevance f∘g
+  right-inverse-of : ∀ x → to (from x) ≡ x
+  right-inverse-of = _≃_.right-inverse-of f∘g
+
+  irrelevance : ∀ y (p : to ⁻¹ y) → (from y , right-inverse-of y) ≡ p
+  irrelevance = _≃_.irrelevance f∘g
hunk ./Equivalence.agda 379
-abstract
-
-  -- Some simplification lemmas.
-
-  right-inverse-of-id :
-    ∀ {a} {A : Set a} {x : A} →
-    _≃_.right-inverse-of id x ≡ refl x
-  right-inverse-of-id {x = x} = refl (refl x)
-
-  left-inverse-of-id :
-    ∀ {a} {A : Set a} {x : A} →
-    _≃_.left-inverse-of id x ≡ refl x
-  left-inverse-of-id {x = x} =
-     left-inverse-of x               ≡⟨⟩
-     left-inverse-of (P.id x)        ≡⟨ sym $ right-left-lemma x ⟩
-     cong P.id (right-inverse-of x)  ≡⟨ sym $ cong-id _ ⟩
-     right-inverse-of x              ≡⟨ right-inverse-of-id ⟩∎
-     refl x                          ∎
-     where open _≃_ id
-
-  right-inverse-of∘inverse :
-    ∀ {a b} {A : Set a} {B : Set b} →
-    ∀ (A≃B : A ≃ B) {x} →
-    _≃_.right-inverse-of (inverse A≃B) x ≡
-    _≃_.left-inverse-of A≃B x
-  right-inverse-of∘inverse A≃B = refl _
-
-  left-inverse-of∘inverse :
-    ∀ {a b} {A : Set a} {B : Set b} →
-    ∀ (A≃B : A ≃ B) {x} →
-    _≃_.left-inverse-of (inverse A≃B) x ≡
-    _≃_.right-inverse-of A≃B x
-  left-inverse-of∘inverse {A = A} {B} A≃B {x} =
-    subst (λ x → _≃_.left-inverse-of (inverse A≃B) x ≡
-                 right-inverse-of x)
-          (right-inverse-of x)
-          (_≃_.left-inverse-of (inverse A≃B) (to (from x))        ≡⟨ sym $ _≃_.right-left-lemma (inverse A≃B) (from x) ⟩
-           cong to (_≃_.right-inverse-of (inverse A≃B) (from x))  ≡⟨ cong (cong to) $ right-inverse-of∘inverse A≃B ⟩
-           cong to (left-inverse-of (from x))                     ≡⟨ left-right-lemma (from x) ⟩∎
-           right-inverse-of (to (from x))                         ∎)
-    where open _≃_ A≃B
+-- Some simplification lemmas.
+
+right-inverse-of-id :
+  ∀ {a} {A : Set a} {x : A} →
+  _≃_.right-inverse-of id x ≡ refl x
+right-inverse-of-id {x = x} = refl (refl x)
+
+left-inverse-of-id :
+  ∀ {a} {A : Set a} {x : A} →
+  _≃_.left-inverse-of id x ≡ refl x
+left-inverse-of-id {x = x} =
+   left-inverse-of x               ≡⟨⟩
+   left-inverse-of (P.id x)        ≡⟨ sym $ right-left-lemma x ⟩
+   cong P.id (right-inverse-of x)  ≡⟨ sym $ cong-id _ ⟩
+   right-inverse-of x              ≡⟨ right-inverse-of-id ⟩∎
+   refl x                          ∎
+   where open _≃_ id
+
+right-inverse-of∘inverse :
+  ∀ {a b} {A : Set a} {B : Set b} →
+  ∀ (A≃B : A ≃ B) {x} →
+  _≃_.right-inverse-of (inverse A≃B) x ≡
+  _≃_.left-inverse-of A≃B x
+right-inverse-of∘inverse A≃B = refl _
+
+left-inverse-of∘inverse :
+  ∀ {a b} {A : Set a} {B : Set b} →
+  ∀ (A≃B : A ≃ B) {x} →
+  _≃_.left-inverse-of (inverse A≃B) x ≡
+  _≃_.right-inverse-of A≃B x
+left-inverse-of∘inverse {A = A} {B} A≃B {x} =
+  subst (λ x → _≃_.left-inverse-of (inverse A≃B) x ≡
+               right-inverse-of x)
+        (right-inverse-of x)
+        (_≃_.left-inverse-of (inverse A≃B) (to (from x))        ≡⟨ sym $ _≃_.right-left-lemma (inverse A≃B) (from x) ⟩
+         cong to (_≃_.right-inverse-of (inverse A≃B) (from x))  ≡⟨ cong (cong to) $ right-inverse-of∘inverse A≃B ⟩
+         cong to (left-inverse-of (from x))                     ≡⟨ left-right-lemma (from x) ⟩∎
+         right-inverse-of (to (from x))                         ∎)
+  where open _≃_ A≃B
hunk ./Equivalence.agda 458
-abstract
-
-  -- Functions between contractible types are equivalences.
-
-  function-between-contractible-types-is-equivalence :
-    ∀ {a b} {A : Set a} {B : Set b} (f : A → B) →
-    Contractible A → Contractible B → Is-equivalence f
-  function-between-contractible-types-is-equivalence f cA cB =
-    Two-out-of-three.g-g∘f
-      (two-out-of-three f (const tt))
-      (lemma cB)
-      (lemma cA)
-    where
-    -- Functions from a contractible type to the unit type are
-    -- contractible.
-
-    lemma : ∀ {b} {C : Set b} → Contractible C →
-            Is-equivalence (λ (_ : C) → tt)
-    lemma (x , irr) _ = (x , refl tt) , λ p →
-      (x , refl tt)  ≡⟨ Σ-≡,≡→≡
-                          (irr (proj₁ p))
-                          (subst (λ _ → tt ≡ tt)
-                             (irr (proj₁ p)) (refl tt)  ≡⟨ elim (λ eq → subst (λ _ → tt ≡ tt) eq (refl tt) ≡ refl tt)
-                                                                (λ _ → subst-refl (λ _ → tt ≡ tt) (refl tt))
-                                                                (irr (proj₁ p)) ⟩
-                           refl tt                      ≡⟨ elim (λ eq → refl tt ≡ eq) (refl ⊚ refl) (proj₂ p) ⟩∎
-                           proj₂ p                      ∎) ⟩∎
-      p              ∎
-
-  -- ext⁻¹ is an equivalence (assuming extensionality).
-
-  ext⁻¹-is-equivalence :
-    ∀ {a b} {A : Set a} →
-    ({B : A → Set b} → Extensionality′ A B) →
-    {B : A → Set b} {f g : (x : A) → B x} →
-    Is-equivalence (ext⁻¹ {f = f} {g = g})
-  ext⁻¹-is-equivalence ext {f = f} {g} =
-    let surj : (∀ x → Singleton (g x)) ↠ (∃ λ f → ∀ x → f x ≡ g x)
-        surj = record
-          { logical-equivalence = record
-            { to   = λ f → proj₁ ⊚ f , proj₂ ⊚ f
-            ; from = λ p x → proj₁ p x , proj₂ p x
-            }
-          ; right-inverse-of = refl
+-- Functions between contractible types are equivalences.
+
+function-between-contractible-types-is-equivalence :
+  ∀ {a b} {A : Set a} {B : Set b} (f : A → B) →
+  Contractible A → Contractible B → Is-equivalence f
+function-between-contractible-types-is-equivalence f cA cB =
+  Two-out-of-three.g-g∘f
+    (two-out-of-three f (const tt))
+    (lemma cB)
+    (lemma cA)
+  where
+  -- Functions from a contractible type to the unit type are
+  -- contractible.
+
+  lemma : ∀ {b} {C : Set b} → Contractible C →
+          Is-equivalence (λ (_ : C) → tt)
+  lemma (x , irr) _ = (x , refl tt) , λ p →
+    (x , refl tt)  ≡⟨ Σ-≡,≡→≡
+                        (irr (proj₁ p))
+                        (subst (λ _ → tt ≡ tt)
+                           (irr (proj₁ p)) (refl tt)  ≡⟨ elim (λ eq → subst (λ _ → tt ≡ tt) eq (refl tt) ≡ refl tt)
+                                                              (λ _ → subst-refl (λ _ → tt ≡ tt) (refl tt))
+                                                              (irr (proj₁ p)) ⟩
+                         refl tt                      ≡⟨ elim (λ eq → refl tt ≡ eq) (refl ⊚ refl) (proj₂ p) ⟩∎
+                         proj₂ p                      ∎) ⟩∎
+    p              ∎
+
+-- ext⁻¹ is an equivalence (assuming extensionality).
+
+ext⁻¹-is-equivalence :
+  ∀ {a b} {A : Set a} →
+  ({B : A → Set b} → Extensionality′ A B) →
+  {B : A → Set b} {f g : (x : A) → B x} →
+  Is-equivalence (ext⁻¹ {f = f} {g = g})
+ext⁻¹-is-equivalence ext {f = f} {g} =
+  let surj : (∀ x → Singleton (g x)) ↠ (∃ λ f → ∀ x → f x ≡ g x)
+      surj = record
+        { logical-equivalence = record
+          { to   = λ f → proj₁ ⊚ f , proj₂ ⊚ f
+          ; from = λ p x → proj₁ p x , proj₂ p x
hunk ./Equivalence.agda 499
-
-        lemma₁ : Contractible (∃ λ f → ∀ x → f x ≡ g x)
-        lemma₁ =
-          H-level.respects-surjection surj 0 $
-            _⇔_.from Π-closure-contractible⇔extensionality
-              ext (singleton-contractible ⊚ g)
-
-        lemma₂ : Is-equivalence (Σ-map P.id ext⁻¹)
-        lemma₂ = function-between-contractible-types-is-equivalence
-                   _ (singleton-contractible g) lemma₁
-
-    in drop-Σ-map-id ext⁻¹ lemma₂ f
+        ; right-inverse-of = refl
+        }
+
+      lemma₁ : Contractible (∃ λ f → ∀ x → f x ≡ g x)
+      lemma₁ =
+        H-level.respects-surjection surj 0 $
+          _⇔_.from Π-closure-contractible⇔extensionality
+            ext (singleton-contractible ⊚ g)
+
+      lemma₂ : Is-equivalence (Σ-map P.id ext⁻¹)
+      lemma₂ = function-between-contractible-types-is-equivalence
+                 _ (singleton-contractible g) lemma₁
+
+  in drop-Σ-map-id ext⁻¹ lemma₂ f
hunk ./Equivalence.agda 530
-abstract
-
-  good-ext-is-equivalence :
-    ∀ {a b} (ext : Extensionality a b) →
-    {A : Set a} {B : A → Set b} {f g : (x : A) → B x} →
-    Is-equivalence {A = ∀ x → f x ≡ g x} (good-ext ext)
-  good-ext-is-equivalence ext =
-    _≃_.is-equivalence (extensionality-isomorphism ext)
-
-  good-ext-refl :
-    ∀ {a b} (ext : Extensionality a b)
-    {A : Set a} {B : A → Set b} (f : (x : A) → B x) →
-    good-ext ext (λ x → refl (f x)) ≡ refl f
-  good-ext-refl ext f =
-    _≃_.to (extensionality-isomorphism ext) (λ x → refl (f x))  ≡⟨ cong (_≃_.to (extensionality-isomorphism ext)) $ sym $
-                                                                        ext (λ _ → ext⁻¹-refl f) ⟩
-    _≃_.to (extensionality-isomorphism ext) (ext⁻¹ (refl f))    ≡⟨ _≃_.right-inverse-of (extensionality-isomorphism ext) _ ⟩∎
-    refl f                                                      ∎
-
-  cong-good-ext :
-    ∀ {a b} (ext : Extensionality a b)
-    {A : Set a} {B : A → Set b} {f g : (x : A) → B x}
-    (f≡g : ∀ x → f x ≡ g x) {x} →
-    cong (λ h → h x) (good-ext ext f≡g) ≡ f≡g x
-  cong-good-ext ext f≡g {x} =
-    let lemma = elim
-          (λ f≡g → cong (λ h → h x) f≡g ≡ ext⁻¹ f≡g x)
-          (λ f → cong (λ h → h x) (refl f)  ≡⟨ cong-refl (λ h → h x) {x = f} ⟩
-                 refl (f x)                 ≡⟨ sym $ ext⁻¹-refl f ⟩∎
-                 ext⁻¹ (refl f) x           ∎)
-          (good-ext ext f≡g)
-    in
-
-    cong (λ h → h x) (good-ext ext f≡g)  ≡⟨ lemma ⟩
-    ext⁻¹ (good-ext ext f≡g) x           ≡⟨ cong (λ h → h x) $
-                                                 _≃_.left-inverse-of (extensionality-isomorphism ext) f≡g ⟩∎
-    f≡g x                                ∎
-
-  subst-good-ext :
-    ∀ {a b p} (ext : Extensionality a b)
-    {A : Set a} {B : A → Set b} {f g : (x : A) → B x} {x}
-    (P : B x → Set p) {p}
-    (f≡g : ∀ x → f x ≡ g x) →
-    subst (λ f → P (f x)) (good-ext ext f≡g) p ≡ subst P (f≡g x) p
-  subst-good-ext ext {f = f} {g} {x = x} P {p} f≡g =
-    subst (λ f → P (f x)) (good-ext ext f≡g) p  ≡⟨ subst-∘ P (_$ x) _ ⟩
-    subst P (cong (_$ x) (good-ext ext f≡g)) p  ≡⟨ cong (λ eq → subst P eq p) (cong-good-ext ext f≡g) ⟩∎
-    subst P (f≡g x) p                           ∎
+good-ext-is-equivalence :
+  ∀ {a b} (ext : Extensionality a b) →
+  {A : Set a} {B : A → Set b} {f g : (x : A) → B x} →
+  Is-equivalence {A = ∀ x → f x ≡ g x} (good-ext ext)
+good-ext-is-equivalence ext =
+  _≃_.is-equivalence (extensionality-isomorphism ext)
+
+good-ext-refl :
+  ∀ {a b} (ext : Extensionality a b)
+  {A : Set a} {B : A → Set b} (f : (x : A) → B x) →
+  good-ext ext (λ x → refl (f x)) ≡ refl f
+good-ext-refl ext f =
+  _≃_.to (extensionality-isomorphism ext) (λ x → refl (f x))  ≡⟨ cong (_≃_.to (extensionality-isomorphism ext)) $ sym $
+                                                                      ext (λ _ → ext⁻¹-refl f) ⟩
+  _≃_.to (extensionality-isomorphism ext) (ext⁻¹ (refl f))    ≡⟨ _≃_.right-inverse-of (extensionality-isomorphism ext) _ ⟩∎
+  refl f                                                      ∎
+
+cong-good-ext :
+  ∀ {a b} (ext : Extensionality a b)
+  {A : Set a} {B : A → Set b} {f g : (x : A) → B x}
+  (f≡g : ∀ x → f x ≡ g x) {x} →
+  cong (λ h → h x) (good-ext ext f≡g) ≡ f≡g x
+cong-good-ext ext f≡g {x} =
+  let lemma = elim
+        (λ f≡g → cong (λ h → h x) f≡g ≡ ext⁻¹ f≡g x)
+        (λ f → cong (λ h → h x) (refl f)  ≡⟨ cong-refl (λ h → h x) {x = f} ⟩
+               refl (f x)                 ≡⟨ sym $ ext⁻¹-refl f ⟩∎
+               ext⁻¹ (refl f) x           ∎)
+        (good-ext ext f≡g)
+  in
+
+  cong (λ h → h x) (good-ext ext f≡g)  ≡⟨ lemma ⟩
+  ext⁻¹ (good-ext ext f≡g) x           ≡⟨ cong (λ h → h x) $
+                                               _≃_.left-inverse-of (extensionality-isomorphism ext) f≡g ⟩∎
+  f≡g x                                ∎
+
+subst-good-ext :
+  ∀ {a b p} (ext : Extensionality a b)
+  {A : Set a} {B : A → Set b} {f g : (x : A) → B x} {x}
+  (P : B x → Set p) {p}
+  (f≡g : ∀ x → f x ≡ g x) →
+  subst (λ f → P (f x)) (good-ext ext f≡g) p ≡ subst P (f≡g x) p
+subst-good-ext ext {f = f} {g} {x = x} P {p} f≡g =
+  subst (λ f → P (f x)) (good-ext ext f≡g) p  ≡⟨ subst-∘ P (_$ x) _ ⟩
+  subst P (cong (_$ x) (good-ext ext f≡g)) p  ≡⟨ cong (λ eq → subst P eq p) (cong-good-ext ext f≡g) ⟩∎
+  subst P (f≡g x) p                           ∎
hunk ./Equivalence.agda 580
-abstract
-
-  -- Two proofs of equivalence are equal if the function components
-  -- are equal (assuming extensionality).
-  --
-  -- See also Function-universe.≃-to-≡↔≡.
-
-  lift-equality :
-    ∀ {a b} → Extensionality (a ⊔ b) (a ⊔ b) →
-    {A : Set a} {B : Set b} {p q : A ≃ B} →
-    _≃_.to p ≡ _≃_.to q → p ≡ q
-  lift-equality {a} {b} ext {p = ⟨ f , f-eq ⟩} {q = ⟨ g , g-eq ⟩} f≡g =
-    elim (λ {f g} f≡g → ∀ f-eq g-eq → ⟨ f , f-eq ⟩ ≡ ⟨ g , g-eq ⟩)
-         (λ f f-eq g-eq →
-            cong (⟨_,_⟩ f)
-              (_⇔_.to {To = Proof-irrelevant _}
-                      propositional⇔irrelevant
-                      (propositional ext f) f-eq g-eq))
-         f≡g f-eq g-eq
-
-  -- Two proofs of equivalence are equal if the /inverses/ of the
-  -- function components are equal (assuming extensionality).
-  --
-  -- See also Function-universe.≃-from-≡↔≡.
-
-  lift-equality-inverse :
-    ∀ {a b} → Extensionality (a ⊔ b) (a ⊔ b) →
-    {A : Set a} {B : Set b} {p q : A ≃ B} →
-    _≃_.from p ≡ _≃_.from q → p ≡ q
-  lift-equality-inverse ext {p = p} {q = q} f≡g =
-    p                    ≡⟨ lift-equality ext (refl _) ⟩
-    inverse (inverse p)  ≡⟨ cong inverse $ lift-equality ext f≡g ⟩
-    inverse (inverse q)  ≡⟨ lift-equality ext (refl _) ⟩∎
-    q                    ∎
+-- Two proofs of equivalence are equal if the function components
+-- are equal (assuming extensionality).
+--
+-- See also Function-universe.≃-to-≡↔≡.
+
+lift-equality :
+  ∀ {a b} → Extensionality (a ⊔ b) (a ⊔ b) →
+  {A : Set a} {B : Set b} {p q : A ≃ B} →
+  _≃_.to p ≡ _≃_.to q → p ≡ q
+lift-equality {a} {b} ext {p = ⟨ f , f-eq ⟩} {q = ⟨ g , g-eq ⟩} f≡g =
+  elim (λ {f g} f≡g → ∀ f-eq g-eq → ⟨ f , f-eq ⟩ ≡ ⟨ g , g-eq ⟩)
+       (λ f f-eq g-eq →
+          cong (⟨_,_⟩ f)
+            (_⇔_.to {To = Proof-irrelevant _}
+                    propositional⇔irrelevant
+                    (propositional ext f) f-eq g-eq))
+       f≡g f-eq g-eq
+
+-- Two proofs of equivalence are equal if the /inverses/ of the
+-- function components are equal (assuming extensionality).
+--
+-- See also Function-universe.≃-from-≡↔≡.
+
+lift-equality-inverse :
+  ∀ {a b} → Extensionality (a ⊔ b) (a ⊔ b) →
+  {A : Set a} {B : Set b} {p q : A ≃ B} →
+  _≃_.from p ≡ _≃_.from q → p ≡ q
+lift-equality-inverse ext {p = p} {q = q} f≡g =
+  p                    ≡⟨ lift-equality ext (refl _) ⟩
+  inverse (inverse p)  ≡⟨ cong inverse $ lift-equality ext f≡g ⟩
+  inverse (inverse q)  ≡⟨ lift-equality ext (refl _) ⟩∎
+  q                    ∎
hunk ./Equivalence.agda 629
-  abstract
-    left-identity : {X Y : Set ℓ} (p : X ≃ Y) → id ∘ p ≡ p
-    left-identity _ = lift-equality ext (refl _)
-
-    right-identity : {X Y : Set ℓ} (p : X ≃ Y) → p ∘ id ≡ p
-    right-identity _ = lift-equality ext (refl _)
-
-    assoc : {W X Y Z : Set ℓ} (p : Y ≃ Z) (q : X ≃ Y) (r : W ≃ X) →
-            p ∘ (q ∘ r) ≡ (p ∘ q) ∘ r
-    assoc _ _ _ = lift-equality ext (refl _)
-
-    left-inverse : {X Y : Set ℓ} (p : X ≃ Y) → inverse p ∘ p ≡ id
-    left-inverse p = lift-equality ext (ext $ _≃_.left-inverse-of p)
-
-    right-inverse : {X Y : Set ℓ} (p : X ≃ Y) → p ∘ inverse p ≡ id
-    right-inverse p = lift-equality ext (ext $ _≃_.right-inverse-of p)
+  left-identity : {X Y : Set ℓ} (p : X ≃ Y) → id ∘ p ≡ p
+  left-identity _ = lift-equality ext (refl _)
+
+  right-identity : {X Y : Set ℓ} (p : X ≃ Y) → p ∘ id ≡ p
+  right-identity _ = lift-equality ext (refl _)
+
+  assoc : {W X Y Z : Set ℓ} (p : Y ≃ Z) (q : X ≃ Y) (r : W ≃ X) →
+          p ∘ (q ∘ r) ≡ (p ∘ q) ∘ r
+  assoc _ _ _ = lift-equality ext (refl _)
+
+  left-inverse : {X Y : Set ℓ} (p : X ≃ Y) → inverse p ∘ p ≡ id
+  left-inverse p = lift-equality ext (ext $ _≃_.left-inverse-of p)
+
+  right-inverse : {X Y : Set ℓ} (p : X ≃ Y) → p ∘ inverse p ≡ id
+  right-inverse p = lift-equality ext (ext $ _≃_.right-inverse-of p)
hunk ./Equivalence.agda 683
- abstract
hunk ./Equivalence.agda 770
-    abstract
-
-      to∘from : ∀ x → to (from x) ≡ x
-      to∘from _ = _⇔_.to propositional⇔irrelevant B-prop _ _
-
-      from∘to : ∀ x → from (to x) ≡ x
-      from∘to _ = _⇔_.to propositional⇔irrelevant A-prop _ _
+    to∘from : ∀ x → to (from x) ≡ x
+    to∘from _ = _⇔_.to propositional⇔irrelevant B-prop _ _
+
+    from∘to : ∀ x → from (to x) ≡ x
+    from∘to _ = _⇔_.to propositional⇔irrelevant A-prop _ _
hunk ./Equivalence.agda 779
-abstract
-
-  -- All h-levels are closed under the equivalence operator (assuming
-  -- extensionality).
-
-  h-level-closure :
-    ∀ {a b} →
-    Extensionality (a ⊔ b) (a ⊔ b) →
-    ∀ {A : Set a} {B : Set b} n →
-    H-level n A → H-level n B → H-level n (A ≃ B)
-  h-level-closure {a} {b} ext {A = A} {B} n hA hB =
-    H-level.respects-surjection
-      (_↔_.surjection $ Bijection.inverse ≃-as-Σ) n lemma₂
-    where
-    lemma₁ : ∀ n {to : A → B} →
-             H-level n A → H-level n B →
-             H-level n (Is-equivalence to)
-    lemma₁ zero    cA cB = sometimes-contractible ext cA (mono₁ 0 cB)
-    lemma₁ (suc n) _  _  = mono (m≤m+n 1 n) (propositional ext _)
-
-    lemma₂ : H-level n (∃ λ (to : A → B) → Is-equivalence to)
-    lemma₂ = Σ-closure n
-              (Π-closure (lower-extensionality b a ext) n (λ _ → hB))
-              (λ _ → lemma₁ n hA hB)
-
-  -- For positive h-levels it is enough if one of the sides has the
-  -- given h-level.
-
-  left-closure :
-    ∀ {a b} →
-    Extensionality (a ⊔ b) (a ⊔ b) →
-    ∀ {A : Set a} {B : Set b} n →
-    H-level (1 + n) A → H-level (1 + n) (A ≃ B)
-  left-closure ext {A = A} {B} n hA =
-    H-level.[inhabited⇒+]⇒+ n λ (A≃B : A ≃ B) →
-      h-level-closure ext (1 + n) hA $
-        H-level.respects-surjection (_≃_.surjection A≃B) (1 + n) hA
-
-  right-closure :
-    ∀ {a b} →
-    Extensionality (a ⊔ b) (a ⊔ b) →
-    ∀ {A : Set a} {B : Set b} n →
-    H-level (1 + n) B → H-level (1 + n) (A ≃ B)
-  right-closure ext {A = A} {B} n hB =
-    H-level.[inhabited⇒+]⇒+ n λ (A≃B : A ≃ B) →
-      left-closure ext n $
-        H-level.respects-surjection
-          (_≃_.surjection (inverse A≃B)) (1 + n) hB
-
-  -- This is not enough for level 0.
-
-  ¬-left-closure :
-    ∀ {a b} →
-    Extensionality (a ⊔ b) (a ⊔ b) →
-    ∃ λ (A : Set a) → ∃ λ (B : Set b) →
-      Contractible A × Is-proposition B × ¬ Contractible (A ≃ B)
-  ¬-left-closure ext =
-    ↑ _ ⊤ ,
-    ⊥ ,
-    ↑-closure 0 ⊤-contractible ,
-    ⊥-propositional ,
-    λ c → ⊥-elim (_≃_.to (proj₁ c) _)
-
-  ¬-right-closure :
-    ∀ {a b} →
-    Extensionality (a ⊔ b) (a ⊔ b) →
-    ∃ λ (A : Set a) → ∃ λ (B : Set b) →
-      Is-proposition A × Contractible B × ¬ Contractible (A ≃ B)
-  ¬-right-closure ext =
-    ⊥ ,
-    ↑ _ ⊤ ,
-    ⊥-propositional ,
-    ↑-closure 0 ⊤-contractible ,
-    λ c → ⊥-elim (_≃_.from (proj₁ c) _)
-
-  -- ⊥ ≃ ⊥ is contractible (assuming extensionality).
-
-  ⊥≃⊥-contractible :
-    ∀ {ℓ₁ ℓ₂} →
-    Extensionality (ℓ₁ ⊔ ℓ₂) (ℓ₁ ⊔ ℓ₂) →
-    Contractible (⊥ {ℓ = ℓ₁} ≃ ⊥ {ℓ = ℓ₂})
-  ⊥≃⊥-contractible {ℓ₁} {ℓ₂} ext =
-    ↔⇒≃ ⊥↔⊥ , λ ⊥↔⊥′ →
-      lift-equality ext $ lower-extensionality ℓ₂ ℓ₁ ext λ x → ⊥-elim x
-    where
-    ⊥↔⊥ : ⊥ {ℓ = ℓ₁} ↔ ⊥ {ℓ = ℓ₂}
-    ⊥↔⊥ = record
-      { surjection = record
-        { logical-equivalence = record
-          { to   = ⊥-elim
-          ; from = ⊥-elim
-          }
-        ; right-inverse-of = λ x → ⊥-elim x
+-- All h-levels are closed under the equivalence operator (assuming
+-- extensionality).
+
+h-level-closure :
+  ∀ {a b} →
+  Extensionality (a ⊔ b) (a ⊔ b) →
+  ∀ {A : Set a} {B : Set b} n →
+  H-level n A → H-level n B → H-level n (A ≃ B)
+h-level-closure {a} {b} ext {A = A} {B} n hA hB =
+  H-level.respects-surjection
+    (_↔_.surjection $ Bijection.inverse ≃-as-Σ) n lemma₂
+  where
+  lemma₁ : ∀ n {to : A → B} →
+           H-level n A → H-level n B →
+           H-level n (Is-equivalence to)
+  lemma₁ zero    cA cB = sometimes-contractible ext cA (mono₁ 0 cB)
+  lemma₁ (suc n) _  _  = mono (m≤m+n 1 n) (propositional ext _)
+
+  lemma₂ : H-level n (∃ λ (to : A → B) → Is-equivalence to)
+  lemma₂ = Σ-closure n
+            (Π-closure (lower-extensionality b a ext) n (λ _ → hB))
+            (λ _ → lemma₁ n hA hB)
+
+-- For positive h-levels it is enough if one of the sides has the
+-- given h-level.
+
+left-closure :
+  ∀ {a b} →
+  Extensionality (a ⊔ b) (a ⊔ b) →
+  ∀ {A : Set a} {B : Set b} n →
+  H-level (1 + n) A → H-level (1 + n) (A ≃ B)
+left-closure ext {A = A} {B} n hA =
+  H-level.[inhabited⇒+]⇒+ n λ (A≃B : A ≃ B) →
+    h-level-closure ext (1 + n) hA $
+      H-level.respects-surjection (_≃_.surjection A≃B) (1 + n) hA
+
+right-closure :
+  ∀ {a b} →
+  Extensionality (a ⊔ b) (a ⊔ b) →
+  ∀ {A : Set a} {B : Set b} n →
+  H-level (1 + n) B → H-level (1 + n) (A ≃ B)
+right-closure ext {A = A} {B} n hB =
+  H-level.[inhabited⇒+]⇒+ n λ (A≃B : A ≃ B) →
+    left-closure ext n $
+      H-level.respects-surjection
+        (_≃_.surjection (inverse A≃B)) (1 + n) hB
+
+-- This is not enough for level 0.
+
+¬-left-closure :
+  ∀ {a b} →
+  Extensionality (a ⊔ b) (a ⊔ b) →
+  ∃ λ (A : Set a) → ∃ λ (B : Set b) →
+    Contractible A × Is-proposition B × ¬ Contractible (A ≃ B)
+¬-left-closure ext =
+  ↑ _ ⊤ ,
+  ⊥ ,
+  ↑-closure 0 ⊤-contractible ,
+  ⊥-propositional ,
+  λ c → ⊥-elim (_≃_.to (proj₁ c) _)
+
+¬-right-closure :
+  ∀ {a b} →
+  Extensionality (a ⊔ b) (a ⊔ b) →
+  ∃ λ (A : Set a) → ∃ λ (B : Set b) →
+    Is-proposition A × Contractible B × ¬ Contractible (A ≃ B)
+¬-right-closure ext =
+  ⊥ ,
+  ↑ _ ⊤ ,
+  ⊥-propositional ,
+  ↑-closure 0 ⊤-contractible ,
+  λ c → ⊥-elim (_≃_.from (proj₁ c) _)
+
+-- ⊥ ≃ ⊥ is contractible (assuming extensionality).
+
+⊥≃⊥-contractible :
+  ∀ {ℓ₁ ℓ₂} →
+  Extensionality (ℓ₁ ⊔ ℓ₂) (ℓ₁ ⊔ ℓ₂) →
+  Contractible (⊥ {ℓ = ℓ₁} ≃ ⊥ {ℓ = ℓ₂})
+⊥≃⊥-contractible {ℓ₁} {ℓ₂} ext =
+  ↔⇒≃ ⊥↔⊥ , λ ⊥↔⊥′ →
+    lift-equality ext $ lower-extensionality ℓ₂ ℓ₁ ext λ x → ⊥-elim x
+  where
+  ⊥↔⊥ : ⊥ {ℓ = ℓ₁} ↔ ⊥ {ℓ = ℓ₂}
+  ⊥↔⊥ = record
+    { surjection = record
+      { logical-equivalence = record
+        { to   = ⊥-elim
+        ; from = ⊥-elim
hunk ./Equivalence.agda 869
-      ; left-inverse-of = λ x → ⊥-elim x
+      ; right-inverse-of = λ x → ⊥-elim x
hunk ./Equivalence.agda 871
+    ; left-inverse-of = λ x → ⊥-elim x
+    }
hunk ./Equivalence.agda 894
-  abstract
-    left-inverse-of′ :
-      ∀ p → _↠_.from surjection′ (_↠_.to surjection′ p) ≡ p
-    left-inverse-of′ = λ to-x≡to-y →
-      cong to (
-        trans (sym (left-inverse-of x)) $
-        trans (cong from to-x≡to-y) $
-        left-inverse-of y)                         ≡⟨ cong-trans to _ _ ⟩
-      trans (cong to (sym (left-inverse-of x))) (
-        cong to (trans (cong from to-x≡to-y) (
-                 left-inverse-of y)))              ≡⟨ cong₂ trans (cong-sym to _) (cong-trans to _ _) ⟩
-      trans (sym (cong to (left-inverse-of x))) (
-        trans (cong to (cong from to-x≡to-y)) (
-        cong to (left-inverse-of y)))              ≡⟨ cong₂ (λ eq₁ eq₂ → trans (sym eq₁) $
-                                                                         trans (cong to (cong from to-x≡to-y)) $
-                                                                         eq₂)
-                                                           (left-right-lemma x)
-                                                           (left-right-lemma y) ⟩
-      trans (sym (right-inverse-of (to x))) (
-        trans (cong to (cong from to-x≡to-y)) (
-        right-inverse-of (to y)))                  ≡⟨ _↠_.right-inverse-of (Surjection.↠-≡ $ _≃_.surjection A≃B) to-x≡to-y ⟩∎
-      to-x≡to-y                                    ∎
-
-abstract
-  private
-
-    -- We can push subst through certain function applications.
-
-    push-subst :
-      ∀ {a₁ a₂ b₁ b₂} {A₁ : Set a₁} {A₂ : Set a₂}
-        (B₁ : A₁ → Set b₁) {B₂ : A₂ → Set b₂}
-        {f : A₂ → A₁} {x₁ x₂ : A₂} {y : B₁ (f x₁)}
-      (g : ∀ x → B₁ (f x) → B₂ x) (eq : x₁ ≡ x₂) →
-      subst B₂ eq (g x₁ y) ≡ g x₂ (subst B₁ (cong f eq) y)
-    push-subst B₁ {B₂} {f} g eq = elim
-      (λ {x₁ x₂} eq → ∀ y → subst B₂ eq (g x₁ y) ≡
-                            g x₂ (subst B₁ (cong f eq) y))
-      (λ x y → subst B₂ (refl x) (g x y)           ≡⟨ subst-refl B₂ _ ⟩
-               g x y                               ≡⟨ sym $ cong (g x) $ subst-refl B₁ _ ⟩
-               g x (subst B₁ (refl (f x)) y)       ≡⟨ cong (λ eq → g x (subst B₁ eq y)) (sym $ cong-refl f) ⟩∎
-               g x (subst B₁ (cong f (refl x)) y)  ∎)
-      eq _
-
-    push-subst′ :
-      ∀ {a₁ a₂ b₁ b₂} {A₁ : Set a₁} {A₂ : Set a₂}
-      (A₁≃A₂ : A₁ ≃ A₂) (B₁ : A₁ → Set b₁) (B₂ : A₂ → Set b₂) →
-      let open _≃_ A₁≃A₂ in {x₁ x₂ : A₁} {y : B₁ (from (to x₁))}
-      (g : ∀ x → B₁ (from (to x)) → B₂ (to x)) (eq : to x₁ ≡ to x₂) →
-      subst B₂ eq (g x₁ y) ≡ g x₂ (subst B₁ (cong from eq) y)
-    push-subst′ A₁≃A₂ B₁ B₂ {x₁} {x₂} {y} g eq =
-      subst B₂ eq (g x₁ y)                    ≡⟨ cong (subst B₂ eq) $ sym $ g′-lemma _ _ ⟩
-      subst B₂ eq (g′ (to x₁) y)              ≡⟨ push-subst B₁ g′ eq ⟩
-      g′ (to x₂) (subst B₁ (cong from eq) y)  ≡⟨ g′-lemma _ _ ⟩∎
-      g x₂ (subst B₁ (cong from eq) y)        ∎
-      where
-      open _≃_ A₁≃A₂
-
-      g′ : ∀ x′ → B₁ (from x′) → B₂ x′
-      g′ x′ y = subst B₂ (right-inverse-of x′) $
-                g (from x′) $
-                subst B₁ (sym $ cong from $ right-inverse-of x′) y
-
-      g′-lemma : ∀ x y → g′ (to x) y ≡ g x y
-      g′-lemma x y =
-        let lemma = λ y →
-              let gy = g (from (to x)) $
-                         subst B₁
-                           (sym $ cong from $ cong to (refl _)) y in
-              subst B₂ (cong to (refl _)) gy             ≡⟨ cong (λ p → subst B₂ p gy) $ cong-refl to ⟩
-              subst B₂ (refl _) gy                       ≡⟨ subst-refl B₂ gy ⟩
-              gy                                         ≡⟨ cong (λ p → g (from (to x)) $ subst B₁ (sym $ cong from p) y) $ cong-refl to ⟩
-              g (from (to x))
-                (subst B₁ (sym $ cong from (refl _)) y)  ≡⟨ cong (λ p → g (from (to x)) $ subst B₁ (sym p) y) $ cong-refl from ⟩
-              g (from (to x))
-                (subst B₁ (sym (refl _)) y)              ≡⟨ cong (λ p → g (from (to x)) $ subst B₁ p y) sym-refl ⟩
-              g (from (to x)) (subst B₁ (refl _) y)      ≡⟨ cong (g (from (to x))) $ subst-refl B₁ y ⟩∎
-              g (from (to x)) y                          ∎
-        in
-        subst B₂ (right-inverse-of (to x))
-          (g (from (to x)) $
-           subst B₁ (sym $ cong from $
-                       right-inverse-of (to x)) y)  ≡⟨ cong (λ p → subst B₂ p (g (from (to x)) $ subst B₁ (sym $ cong from p) y)) $
-                                                         sym $ left-right-lemma x ⟩
-        subst B₂ (cong to $ left-inverse-of x)
-          (g (from (to x)) $
-           subst B₁ (sym $ cong from $ cong to $
-                       left-inverse-of x) y)        ≡⟨ elim¹
-                                                         (λ {x′} eq →
-                                                            (y : B₁ (from (to x′))) →
-                                                            subst B₂ (cong to eq)
-                                                              (g (from (to x)) $ subst B₁ (sym $ cong from $ cong to eq) y) ≡
-                                                            g x′ y)
-                                                         lemma
-                                                         (left-inverse-of x) y ⟩∎
-        g x y                                       ∎
+  left-inverse-of′ :
+    ∀ p → _↠_.from surjection′ (_↠_.to surjection′ p) ≡ p
+  left-inverse-of′ = λ to-x≡to-y →
+    cong to (
+      trans (sym (left-inverse-of x)) $
+      trans (cong from to-x≡to-y) $
+      left-inverse-of y)                         ≡⟨ cong-trans to _ _ ⟩
+    trans (cong to (sym (left-inverse-of x))) (
+      cong to (trans (cong from to-x≡to-y) (
+               left-inverse-of y)))              ≡⟨ cong₂ trans (cong-sym to _) (cong-trans to _ _) ⟩
+    trans (sym (cong to (left-inverse-of x))) (
+      trans (cong to (cong from to-x≡to-y)) (
+      cong to (left-inverse-of y)))              ≡⟨ cong₂ (λ eq₁ eq₂ → trans (sym eq₁) $
+                                                                       trans (cong to (cong from to-x≡to-y)) $
+                                                                       eq₂)
+                                                         (left-right-lemma x)
+                                                         (left-right-lemma y) ⟩
+    trans (sym (right-inverse-of (to x))) (
+      trans (cong to (cong from to-x≡to-y)) (
+      right-inverse-of (to y)))                  ≡⟨ _↠_.right-inverse-of (Surjection.↠-≡ $ _≃_.surjection A≃B) to-x≡to-y ⟩∎
+    to-x≡to-y                                    ∎
+
+private
+
+  -- We can push subst through certain function applications.
+
+  push-subst :
+    ∀ {a₁ a₂ b₁ b₂} {A₁ : Set a₁} {A₂ : Set a₂}
+      (B₁ : A₁ → Set b₁) {B₂ : A₂ → Set b₂}
+      {f : A₂ → A₁} {x₁ x₂ : A₂} {y : B₁ (f x₁)}
+    (g : ∀ x → B₁ (f x) → B₂ x) (eq : x₁ ≡ x₂) →
+    subst B₂ eq (g x₁ y) ≡ g x₂ (subst B₁ (cong f eq) y)
+  push-subst B₁ {B₂} {f} g eq = elim
+    (λ {x₁ x₂} eq → ∀ y → subst B₂ eq (g x₁ y) ≡
+                          g x₂ (subst B₁ (cong f eq) y))
+    (λ x y → subst B₂ (refl x) (g x y)           ≡⟨ subst-refl B₂ _ ⟩
+             g x y                               ≡⟨ sym $ cong (g x) $ subst-refl B₁ _ ⟩
+             g x (subst B₁ (refl (f x)) y)       ≡⟨ cong (λ eq → g x (subst B₁ eq y)) (sym $ cong-refl f) ⟩∎
+             g x (subst B₁ (cong f (refl x)) y)  ∎)
+    eq _
+
+  push-subst′ :
+    ∀ {a₁ a₂ b₁ b₂} {A₁ : Set a₁} {A₂ : Set a₂}
+    (A₁≃A₂ : A₁ ≃ A₂) (B₁ : A₁ → Set b₁) (B₂ : A₂ → Set b₂) →
+    let open _≃_ A₁≃A₂ in {x₁ x₂ : A₁} {y : B₁ (from (to x₁))}
+    (g : ∀ x → B₁ (from (to x)) → B₂ (to x)) (eq : to x₁ ≡ to x₂) →
+    subst B₂ eq (g x₁ y) ≡ g x₂ (subst B₁ (cong from eq) y)
+  push-subst′ A₁≃A₂ B₁ B₂ {x₁} {x₂} {y} g eq =
+    subst B₂ eq (g x₁ y)                    ≡⟨ cong (subst B₂ eq) $ sym $ g′-lemma _ _ ⟩
+    subst B₂ eq (g′ (to x₁) y)              ≡⟨ push-subst B₁ g′ eq ⟩
+    g′ (to x₂) (subst B₁ (cong from eq) y)  ≡⟨ g′-lemma _ _ ⟩∎
+    g x₂ (subst B₁ (cong from eq) y)        ∎
+    where
+    open _≃_ A₁≃A₂
+
+    g′ : ∀ x′ → B₁ (from x′) → B₂ x′
+    g′ x′ y = subst B₂ (right-inverse-of x′) $
+              g (from x′) $
+              subst B₁ (sym $ cong from $ right-inverse-of x′) y
+
+    g′-lemma : ∀ x y → g′ (to x) y ≡ g x y
+    g′-lemma x y =
+      let lemma = λ y →
+            let gy = g (from (to x)) $
+                       subst B₁
+                         (sym $ cong from $ cong to (refl _)) y in
+            subst B₂ (cong to (refl _)) gy             ≡⟨ cong (λ p → subst B₂ p gy) $ cong-refl to ⟩
+            subst B₂ (refl _) gy                       ≡⟨ subst-refl B₂ gy ⟩
+            gy                                         ≡⟨ cong (λ p → g (from (to x)) $ subst B₁ (sym $ cong from p) y) $ cong-refl to ⟩
+            g (from (to x))
+              (subst B₁ (sym $ cong from (refl _)) y)  ≡⟨ cong (λ p → g (from (to x)) $ subst B₁ (sym p) y) $ cong-refl from ⟩
+            g (from (to x))
+              (subst B₁ (sym (refl _)) y)              ≡⟨ cong (λ p → g (from (to x)) $ subst B₁ p y) sym-refl ⟩
+            g (from (to x)) (subst B₁ (refl _) y)      ≡⟨ cong (g (from (to x))) $ subst-refl B₁ y ⟩∎
+            g (from (to x)) y                          ∎
+      in
+      subst B₂ (right-inverse-of (to x))
+        (g (from (to x)) $
+         subst B₁ (sym $ cong from $
+                     right-inverse-of (to x)) y)  ≡⟨ cong (λ p → subst B₂ p (g (from (to x)) $ subst B₁ (sym $ cong from p) y)) $
+                                                       sym $ left-right-lemma x ⟩
+      subst B₂ (cong to $ left-inverse-of x)
+        (g (from (to x)) $
+         subst B₁ (sym $ cong from $ cong to $
+                     left-inverse-of x) y)        ≡⟨ elim¹
+                                                       (λ {x′} eq →
+                                                          (y : B₁ (from (to x′))) →
+                                                          subst B₂ (cong to eq)
+                                                            (g (from (to x)) $ subst B₁ (sym $ cong from $ cong to eq) y) ≡
+                                                          g x′ y)
+                                                       lemma
+                                                       (left-inverse-of x) y ⟩∎
+      g x y                                       ∎
hunk ./Equivalence.agda 1029
-  abstract
-    injective′ : Injective to′
-    injective′ {x = (x₁ , x₂)} {y = (y₁ , y₂)} =
-      _↔_.to Σ-≡,≡↔≡ ⊚
-      Σ-map (_≃_.injective A₁≃A₂) (λ {eq₁} eq₂ →
-        let lemma =
-              to (B₁↣B₂ y₁) (subst B₁ (_≃_.injective A₁≃A₂ eq₁) x₂)      ≡⟨ refl _ ⟩
-              to (B₁↣B₂ y₁)
-                (subst B₁ (trans (sym (_≃_.left-inverse-of A₁≃A₂ x₁)) $
-                           trans (cong (_≃_.from A₁≃A₂) eq₁)
-                                 (_≃_.left-inverse-of A₁≃A₂ y₁))
-                          x₂)                                            ≡⟨ cong (to (B₁↣B₂ y₁)) $ sym $ subst-subst B₁ _ _ _ ⟩
-              to (B₁↣B₂ y₁)
-                 (subst B₁ (trans (cong (_≃_.from A₁≃A₂) eq₁)
-                                  (_≃_.left-inverse-of A₁≃A₂ y₁)) $
-                  subst B₁ (sym (_≃_.left-inverse-of A₁≃A₂ x₁))
-                           x₂)                                           ≡⟨ cong (to (B₁↣B₂ y₁)) $ sym $ subst-subst B₁ _ _ _ ⟩
-              to (B₁↣B₂ y₁)
-                 (subst B₁ (_≃_.left-inverse-of A₁≃A₂ y₁) $
-                  subst B₁ (cong (_≃_.from A₁≃A₂) eq₁) $
-                  subst B₁ (sym (_≃_.left-inverse-of A₁≃A₂ x₁)) x₂)      ≡⟨ sym $ push-subst′ A₁≃A₂ B₁ B₂
-                                                                              (λ x y → to (B₁↣B₂ x)
-                                                                                          (subst B₁ (_≃_.left-inverse-of A₁≃A₂ x) y))
-                                                                              eq₁ ⟩
-              subst B₂ eq₁
-                (to (B₁↣B₂ x₁) $
-                 subst B₁ (_≃_.left-inverse-of A₁≃A₂ x₁) $
-                 subst B₁ (sym (_≃_.left-inverse-of A₁≃A₂ x₁)) x₂)       ≡⟨ cong (subst B₂ eq₁ ⊚ to (B₁↣B₂ x₁)) $ subst-subst B₁ _ _ _ ⟩
-              subst B₂ eq₁
-                (to (B₁↣B₂ x₁) $
-                 subst B₁ (trans (sym (_≃_.left-inverse-of A₁≃A₂ x₁))
-                                 (_≃_.left-inverse-of A₁≃A₂ x₁))
-                          x₂)                                            ≡⟨ cong (λ p → subst B₂ eq₁ (to (B₁↣B₂ x₁) (subst B₁ p x₂))) $
-                                                                                 trans-symˡ _ ⟩
-              subst B₂ eq₁ (to (B₁↣B₂ x₁) $ subst B₁ (refl _) x₂)        ≡⟨ cong (subst B₂ eq₁ ⊚ to (B₁↣B₂ x₁)) $
-                                                                                 subst-refl B₁ x₂ ⟩
-              subst B₂ eq₁ (to (B₁↣B₂ x₁) x₂)                            ≡⟨ eq₂ ⟩∎
-              to (B₁↣B₂ y₁) y₂                                           ∎
-        in
-        subst B₁ (_≃_.injective A₁≃A₂ eq₁) x₂  ≡⟨ _↣_.injective (B₁↣B₂ y₁) lemma ⟩∎
-        y₂                                     ∎) ⊚
-      Σ-≡,≡←≡
+  injective′ : Injective to′
+  injective′ {x = (x₁ , x₂)} {y = (y₁ , y₂)} =
+    _↔_.to Σ-≡,≡↔≡ ⊚
+    Σ-map (_≃_.injective A₁≃A₂) (λ {eq₁} eq₂ →
+      let lemma =
+            to (B₁↣B₂ y₁) (subst B₁ (_≃_.injective A₁≃A₂ eq₁) x₂)      ≡⟨ refl _ ⟩
+            to (B₁↣B₂ y₁)
+              (subst B₁ (trans (sym (_≃_.left-inverse-of A₁≃A₂ x₁)) $
+                         trans (cong (_≃_.from A₁≃A₂) eq₁)
+                               (_≃_.left-inverse-of A₁≃A₂ y₁))
+                        x₂)                                            ≡⟨ cong (to (B₁↣B₂ y₁)) $ sym $ subst-subst B₁ _ _ _ ⟩
+            to (B₁↣B₂ y₁)
+               (subst B₁ (trans (cong (_≃_.from A₁≃A₂) eq₁)
+                                (_≃_.left-inverse-of A₁≃A₂ y₁)) $
+                subst B₁ (sym (_≃_.left-inverse-of A₁≃A₂ x₁))
+                         x₂)                                           ≡⟨ cong (to (B₁↣B₂ y₁)) $ sym $ subst-subst B₁ _ _ _ ⟩
+            to (B₁↣B₂ y₁)
+               (subst B₁ (_≃_.left-inverse-of A₁≃A₂ y₁) $
+                subst B₁ (cong (_≃_.from A₁≃A₂) eq₁) $
+                subst B₁ (sym (_≃_.left-inverse-of A₁≃A₂ x₁)) x₂)      ≡⟨ sym $ push-subst′ A₁≃A₂ B₁ B₂
+                                                                            (λ x y → to (B₁↣B₂ x)
+                                                                                        (subst B₁ (_≃_.left-inverse-of A₁≃A₂ x) y))
+                                                                            eq₁ ⟩
+            subst B₂ eq₁
+              (to (B₁↣B₂ x₁) $
+               subst B₁ (_≃_.left-inverse-of A₁≃A₂ x₁) $
+               subst B₁ (sym (_≃_.left-inverse-of A₁≃A₂ x₁)) x₂)       ≡⟨ cong (subst B₂ eq₁ ⊚ to (B₁↣B₂ x₁)) $ subst-subst B₁ _ _ _ ⟩
+            subst B₂ eq₁
+              (to (B₁↣B₂ x₁) $
+               subst B₁ (trans (sym (_≃_.left-inverse-of A₁≃A₂ x₁))
+                               (_≃_.left-inverse-of A₁≃A₂ x₁))
+                        x₂)                                            ≡⟨ cong (λ p → subst B₂ eq₁ (to (B₁↣B₂ x₁) (subst B₁ p x₂))) $
+                                                                               trans-symˡ _ ⟩
+            subst B₂ eq₁ (to (B₁↣B₂ x₁) $ subst B₁ (refl _) x₂)        ≡⟨ cong (subst B₂ eq₁ ⊚ to (B₁↣B₂ x₁)) $
+                                                                               subst-refl B₁ x₂ ⟩
+            subst B₂ eq₁ (to (B₁↣B₂ x₁) x₂)                            ≡⟨ eq₂ ⟩∎
+            to (B₁↣B₂ y₁) y₂                                           ∎
+      in
+      subst B₁ (_≃_.injective A₁≃A₂ eq₁) x₂  ≡⟨ _↣_.injective (B₁↣B₂ y₁) lemma ⟩∎
+      y₂                                     ∎) ⊚
+    Σ-≡,≡←≡
hunk ./Equivalence.agda 1087
-  abstract
-    right-inverse-of′ :
-      ∀ p →
-      _⇔_.to logical-equivalence′ (_⇔_.from logical-equivalence′ p) ≡ p
-    right-inverse-of′ = λ p → Σ-≡,≡→≡
-      (_≃_.right-inverse-of A₁≃A₂ (proj₁ p))
-      (subst B₂ (_≃_.right-inverse-of A₁≃A₂ (proj₁ p))
-         (to (B₁↠B₂ _) (from (B₁↠B₂ _)
-            (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂ (proj₁ p)))
-               (proj₂ p))))                                         ≡⟨ cong (subst B₂ _) $ right-inverse-of (B₁↠B₂ _) _ ⟩
-       subst B₂ (_≃_.right-inverse-of A₁≃A₂ (proj₁ p))
-         (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂ (proj₁ p)))
-            (proj₂ p))                                              ≡⟨ subst-subst-sym B₂ _ _ ⟩∎
-       proj₂ p ∎)
+  right-inverse-of′ :
+    ∀ p →
+    _⇔_.to logical-equivalence′ (_⇔_.from logical-equivalence′ p) ≡ p
+  right-inverse-of′ = λ p → Σ-≡,≡→≡
+    (_≃_.right-inverse-of A₁≃A₂ (proj₁ p))
+    (subst B₂ (_≃_.right-inverse-of A₁≃A₂ (proj₁ p))
+       (to (B₁↠B₂ _) (from (B₁↠B₂ _)
+          (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂ (proj₁ p)))
+             (proj₂ p))))                                         ≡⟨ cong (subst B₂ _) $ right-inverse-of (B₁↠B₂ _) _ ⟩
+     subst B₂ (_≃_.right-inverse-of A₁≃A₂ (proj₁ p))
+       (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂ (proj₁ p)))
+          (proj₂ p))                                              ≡⟨ subst-subst-sym B₂ _ _ ⟩∎
+     proj₂ p ∎)
hunk ./Equivalence.agda 1116
-  abstract
-    left-inverse-of′ :
-      ∀ p → _↠_.from surjection′ (_↠_.to surjection′ p) ≡ p
-    left-inverse-of′ = λ p → Σ-≡,≡→≡
-      (_≃_.left-inverse-of A₁≃A₂ (proj₁ p))
-      (subst B₁ (_≃_.left-inverse-of A₁≃A₂ (proj₁ p))
-         (from (B₁↔B₂ _)
-            (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂
-                              (_≃_.to A₁≃A₂ (proj₁ p))))
-               (to (B₁↔B₂ _) (proj₂ p))))                        ≡⟨ push-subst B₂ (λ x → from (B₁↔B₂ x))
-                                                                      (_≃_.left-inverse-of A₁≃A₂ (proj₁ p)) ⟩
-       from (B₁↔B₂ _)
-         (subst B₂ (cong (_≃_.to A₁≃A₂)
-                         (_≃_.left-inverse-of A₁≃A₂ (proj₁ p)))
-            (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂
-                              (_≃_.to A₁≃A₂ (proj₁ p))))
-               (to (B₁↔B₂ _) (proj₂ p))))                        ≡⟨ cong (λ eq → from (B₁↔B₂ _)
-                                                                                   (subst B₂ eq
-                                                                                      (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂ _))
-                                                                                         (to (B₁↔B₂ _) (proj₂ p))))) $
-                                                                         _≃_.left-right-lemma A₁≃A₂ _ ⟩
-       from (B₁↔B₂ _)
-         (subst B₂ (_≃_.right-inverse-of A₁≃A₂
-                      (_≃_.to A₁≃A₂ (proj₁ p)))
-            (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂
-                              (_≃_.to A₁≃A₂ (proj₁ p))))
-               (to (B₁↔B₂ _) (proj₂ p))))                        ≡⟨ cong (from (B₁↔B₂ _)) $ subst-subst-sym B₂ _ _ ⟩
-       from (B₁↔B₂ _) (to (B₁↔B₂ _) (proj₂ p))                   ≡⟨ left-inverse-of (B₁↔B₂ _) _ ⟩∎
-       proj₂ p                                                   ∎)
+  left-inverse-of′ :
+    ∀ p → _↠_.from surjection′ (_↠_.to surjection′ p) ≡ p
+  left-inverse-of′ = λ p → Σ-≡,≡→≡
+    (_≃_.left-inverse-of A₁≃A₂ (proj₁ p))
+    (subst B₁ (_≃_.left-inverse-of A₁≃A₂ (proj₁ p))
+       (from (B₁↔B₂ _)
+          (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂
+                            (_≃_.to A₁≃A₂ (proj₁ p))))
+             (to (B₁↔B₂ _) (proj₂ p))))                        ≡⟨ push-subst B₂ (λ x → from (B₁↔B₂ x))
+                                                                    (_≃_.left-inverse-of A₁≃A₂ (proj₁ p)) ⟩
+     from (B₁↔B₂ _)
+       (subst B₂ (cong (_≃_.to A₁≃A₂)
+                       (_≃_.left-inverse-of A₁≃A₂ (proj₁ p)))
+          (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂
+                            (_≃_.to A₁≃A₂ (proj₁ p))))
+             (to (B₁↔B₂ _) (proj₂ p))))                        ≡⟨ cong (λ eq → from (B₁↔B₂ _)
+                                                                                 (subst B₂ eq
+                                                                                    (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂ _))
+                                                                                       (to (B₁↔B₂ _) (proj₂ p))))) $
+                                                                       _≃_.left-right-lemma A₁≃A₂ _ ⟩
+     from (B₁↔B₂ _)
+       (subst B₂ (_≃_.right-inverse-of A₁≃A₂
+                    (_≃_.to A₁≃A₂ (proj₁ p)))
+          (subst B₂ (sym (_≃_.right-inverse-of A₁≃A₂
+                            (_≃_.to A₁≃A₂ (proj₁ p))))
+             (to (B₁↔B₂ _) (proj₂ p))))                        ≡⟨ cong (from (B₁↔B₂ _)) $ subst-subst-sym B₂ _ _ ⟩
+     from (B₁↔B₂ _) (to (B₁↔B₂ _) (proj₂ p))                   ≡⟨ left-inverse-of (B₁↔B₂ _) _ ⟩∎
+     proj₂ p                                                   ∎)
hunk ./Equivalence.agda 1183
-  abstract
-    right-inverse-of′ : ∀ f → to′ (from′ f) ≡ f
-    right-inverse-of′ = λ f → lower-extensionality a₁ b₁ ext λ x →
-      subst B₂ (right-inverse-of A₁≃A₂ x)
-            (to (B₁≃B₂ (from A₁≃A₂ x))
-                (from (B₁≃B₂ (from A₁≃A₂ x))
-                      (f (to A₁≃A₂ (from A₁≃A₂ x)))))  ≡⟨ cong (subst B₂ (right-inverse-of A₁≃A₂ x)) $
-                                                               right-inverse-of (B₁≃B₂ _) _ ⟩
-      subst B₂ (right-inverse-of A₁≃A₂ x)
-            (f (to A₁≃A₂ (from A₁≃A₂ x)))              ≡⟨ elim (λ {x y} x≡y → subst B₂ x≡y (f x) ≡ f y)
-                                                               (λ x → subst-refl B₂ (f x))
-                                                               (right-inverse-of A₁≃A₂ x) ⟩∎
-      f x                                              ∎
-
-    left-inverse-of′ : ∀ f → from′ (to′ f) ≡ f
-    left-inverse-of′ = λ f → lower-extensionality a₂ b₂ ext λ x →
-      from (B₁≃B₂ x)
-           (subst B₂ (right-inverse-of A₁≃A₂ (to A₁≃A₂ x))
-                  (to (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
-                      (f (from A₁≃A₂ (to A₁≃A₂ x)))))             ≡⟨ cong (λ eq → from (B₁≃B₂ x)
-                                                                                    (subst B₂ eq
-                                                                                       (to (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
-                                                                                          (f (from A₁≃A₂ (to A₁≃A₂ x))))))
-                                                                          (sym $ left-right-lemma A₁≃A₂ x) ⟩
-      from (B₁≃B₂ x)
-           (subst B₂ (cong (to A₁≃A₂) (left-inverse-of A₁≃A₂ x))
-                  (to (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
-                      (f (from A₁≃A₂ (to A₁≃A₂ x)))))             ≡⟨ sym $ push-subst B₂ (λ x y → from (B₁≃B₂ x) y)
-                                                                                      (left-inverse-of A₁≃A₂ x) ⟩
-      subst B₁ (left-inverse-of A₁≃A₂ x)
-            (from (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
-                  (to (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
-                      (f (from A₁≃A₂ (to A₁≃A₂ x)))))             ≡⟨ cong (subst B₁ (left-inverse-of A₁≃A₂ x)) $
-                                                                       left-inverse-of (B₁≃B₂ _) _ ⟩
-      subst B₁ (left-inverse-of A₁≃A₂ x)
-            (f (from A₁≃A₂ (to A₁≃A₂ x)))                         ≡⟨ elim (λ {x y} x≡y → subst B₁ x≡y (f x) ≡ f y)
-                                                                          (λ x → subst-refl B₁ (f x))
-                                                                          (left-inverse-of A₁≃A₂ x) ⟩∎
-      f x                                                         ∎
+  right-inverse-of′ : ∀ f → to′ (from′ f) ≡ f
+  right-inverse-of′ = λ f → lower-extensionality a₁ b₁ ext λ x →
+    subst B₂ (right-inverse-of A₁≃A₂ x)
+          (to (B₁≃B₂ (from A₁≃A₂ x))
+              (from (B₁≃B₂ (from A₁≃A₂ x))
+                    (f (to A₁≃A₂ (from A₁≃A₂ x)))))  ≡⟨ cong (subst B₂ (right-inverse-of A₁≃A₂ x)) $
+                                                             right-inverse-of (B₁≃B₂ _) _ ⟩
+    subst B₂ (right-inverse-of A₁≃A₂ x)
+          (f (to A₁≃A₂ (from A₁≃A₂ x)))              ≡⟨ elim (λ {x y} x≡y → subst B₂ x≡y (f x) ≡ f y)
+                                                             (λ x → subst-refl B₂ (f x))
+                                                             (right-inverse-of A₁≃A₂ x) ⟩∎
+    f x                                              ∎
+
+  left-inverse-of′ : ∀ f → from′ (to′ f) ≡ f
+  left-inverse-of′ = λ f → lower-extensionality a₂ b₂ ext λ x →
+    from (B₁≃B₂ x)
+         (subst B₂ (right-inverse-of A₁≃A₂ (to A₁≃A₂ x))
+                (to (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
+                    (f (from A₁≃A₂ (to A₁≃A₂ x)))))             ≡⟨ cong (λ eq → from (B₁≃B₂ x)
+                                                                                  (subst B₂ eq
+                                                                                     (to (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
+                                                                                        (f (from A₁≃A₂ (to A₁≃A₂ x))))))
+                                                                        (sym $ left-right-lemma A₁≃A₂ x) ⟩
+    from (B₁≃B₂ x)
+         (subst B₂ (cong (to A₁≃A₂) (left-inverse-of A₁≃A₂ x))
+                (to (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
+                    (f (from A₁≃A₂ (to A₁≃A₂ x)))))             ≡⟨ sym $ push-subst B₂ (λ x y → from (B₁≃B₂ x) y)
+                                                                                    (left-inverse-of A₁≃A₂ x) ⟩
+    subst B₁ (left-inverse-of A₁≃A₂ x)
+          (from (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
+                (to (B₁≃B₂ (from A₁≃A₂ (to A₁≃A₂ x)))
+                    (f (from A₁≃A₂ (to A₁≃A₂ x)))))             ≡⟨ cong (subst B₁ (left-inverse-of A₁≃A₂ x)) $
+                                                                     left-inverse-of (B₁≃B₂ _) _ ⟩
+    subst B₁ (left-inverse-of A₁≃A₂ x)
+          (f (from A₁≃A₂ (to A₁≃A₂ x)))                         ≡⟨ elim (λ {x y} x≡y → subst B₁ x≡y (f x) ≡ f y)
+                                                                        (λ x → subst-refl B₁ (f x))
+                                                                        (left-inverse-of A₁≃A₂ x) ⟩∎
+    f x                                                         ∎
hunk ./Equivalence.agda 1252
-  abstract
-    right-inverse-of′ : ∀ f → to′ (from′ f) ≡ f
-    right-inverse-of′ = λ f → lower-extensionality a b₁ ext λ x →
-      to (B₁≃B₂ x) (from (B₁≃B₂ x) (f x))  ≡⟨ right-inverse-of (B₁≃B₂ x) (f x) ⟩∎
-      f x                                  ∎
-
-    left-inverse-of′ : ∀ f → from′ (to′ f) ≡ f
-    left-inverse-of′ = λ f → lower-extensionality a b₂ ext λ x →
-      from (B₁≃B₂ x) (to (B₁≃B₂ x) (f x))  ≡⟨ left-inverse-of (B₁≃B₂ x) (f x) ⟩∎
-      f x                                  ∎
+  right-inverse-of′ : ∀ f → to′ (from′ f) ≡ f
+  right-inverse-of′ = λ f → lower-extensionality a b₁ ext λ x →
+    to (B₁≃B₂ x) (from (B₁≃B₂ x) (f x))  ≡⟨ right-inverse-of (B₁≃B₂ x) (f x) ⟩∎
+    f x                                  ∎
+
+  left-inverse-of′ : ∀ f → from′ (to′ f) ≡ f
+  left-inverse-of′ = λ f → lower-extensionality a b₂ ext λ x →
+    from (B₁≃B₂ x) (to (B₁≃B₂ x) (f x))  ≡⟨ left-inverse-of (B₁≃B₂ x) (f x) ⟩∎
+    f x                                  ∎
hunk ./Equivalence.agda 1283
-  abstract
-    to∘from :
-      (A₂≃B₂ : A₂ ≃ B₂) →
-      B₁≃B₂ ∘ (inverse B₁≃B₂ ∘ A₂≃B₂ ∘ A₁≃A₂) ∘ inverse A₁≃A₂ ≡ A₂≃B₂
-    to∘from A₂≃B₂ =
-      lift-equality (lower-extensionality (a₁ ⊔ b₁) (a₁ ⊔ b₁) ext) $
-        lower-extensionality (a₁ ⊔ b₁ ⊔ b₂) (a₁ ⊔ a₂ ⊔ b₁) ext λ x →
-          to B₁≃B₂ (from B₁≃B₂ (to A₂≃B₂ (to A₁≃A₂ (from A₁≃A₂ x))))  ≡⟨ right-inverse-of B₁≃B₂ _ ⟩
-          to A₂≃B₂ (to A₁≃A₂ (from A₁≃A₂ x))                          ≡⟨ cong (to A₂≃B₂) $ right-inverse-of A₁≃A₂ _ ⟩∎
-          to A₂≃B₂ x                                                  ∎
-
-    from∘to :
-      (A₁≃B₁ : A₁ ≃ B₁) →
-      inverse B₁≃B₂ ∘ (B₁≃B₂ ∘ A₁≃B₁ ∘ inverse A₁≃A₂) ∘ A₁≃A₂ ≡ A₁≃B₁
-    from∘to A₁≃B₁ =
-      lift-equality (lower-extensionality (a₂ ⊔ b₂) (a₂ ⊔ b₂) ext) $
-        lower-extensionality (a₂ ⊔ b₁ ⊔ b₂) (a₁ ⊔ a₂ ⊔ b₂) ext λ x →
-          from B₁≃B₂ (to B₁≃B₂ (to A₁≃B₁ (from A₁≃A₂ (to A₁≃A₂ x))))  ≡⟨ left-inverse-of B₁≃B₂ _ ⟩
-          to A₁≃B₁ (from A₁≃A₂ (to A₁≃A₂ x))                          ≡⟨ cong (to A₁≃B₁) $ left-inverse-of A₁≃A₂ _ ⟩∎
-          to A₁≃B₁ x                                                  ∎
+  to∘from :
+    (A₂≃B₂ : A₂ ≃ B₂) →
+    B₁≃B₂ ∘ (inverse B₁≃B₂ ∘ A₂≃B₂ ∘ A₁≃A₂) ∘ inverse A₁≃A₂ ≡ A₂≃B₂
+  to∘from A₂≃B₂ =
+    lift-equality (lower-extensionality (a₁ ⊔ b₁) (a₁ ⊔ b₁) ext) $
+      lower-extensionality (a₁ ⊔ b₁ ⊔ b₂) (a₁ ⊔ a₂ ⊔ b₁) ext λ x →
+        to B₁≃B₂ (from B₁≃B₂ (to A₂≃B₂ (to A₁≃A₂ (from A₁≃A₂ x))))  ≡⟨ right-inverse-of B₁≃B₂ _ ⟩
+        to A₂≃B₂ (to A₁≃A₂ (from A₁≃A₂ x))                          ≡⟨ cong (to A₂≃B₂) $ right-inverse-of A₁≃A₂ _ ⟩∎
+        to A₂≃B₂ x                                                  ∎
+
+  from∘to :
+    (A₁≃B₁ : A₁ ≃ B₁) →
+    inverse B₁≃B₂ ∘ (B₁≃B₂ ∘ A₁≃B₁ ∘ inverse A₁≃A₂) ∘ A₁≃A₂ ≡ A₁≃B₁
+  from∘to A₁≃B₁ =
+    lift-equality (lower-extensionality (a₂ ⊔ b₂) (a₂ ⊔ b₂) ext) $
+      lower-extensionality (a₂ ⊔ b₁ ⊔ b₂) (a₁ ⊔ a₂ ⊔ b₂) ext λ x →
+        from B₁≃B₂ (to B₁≃B₂ (to A₁≃B₁ (from A₁≃A₂ (to A₁≃A₂ x))))  ≡⟨ left-inverse-of B₁≃B₂ _ ⟩
+        to A₁≃B₁ (from A₁≃A₂ (to A₁≃A₂ x))                          ≡⟨ cong (to A₁≃B₁) $ left-inverse-of A₁≃A₂ _ ⟩∎
+        to A₁≃B₁ x                                                  ∎
hunk ./Equivalence.agda 1316
-abstract
-
-  -- As a consequence of extensionality-isomorphism and ≃-≡ we get a
-  -- strengthening of W-≡,≡↠≡.
-
-  W-≡,≡≃≡ : ∀ {a b} {A : Set a} {B : A → Set b} →
-            (∀ {x} {C : B x → Set (a ⊔ b)} → Extensionality′ (B x) C) →
-            ∀ {x y} {f : B x → W A B} {g : B y → W A B} →
-            (∃ λ (p : x ≡ y) → ∀ i → f i ≡ g (subst B p i)) ≃
-            (sup x f ≡ sup y g)
-  W-≡,≡≃≡ {a} {A = A} {B} ext {x} {y} {f} {g} =
-    (∃ λ p → ∀ i → f i ≡ g (subst B p i))        ≃⟨ Σ-preserves id lemma ⟩
-    (∃ λ p → subst (λ x → B x → W A B) p f ≡ g)  ≃⟨ ↔⇒≃ Σ-≡,≡↔≡ ⟩
-    ((x , f) ≡ (y , g))                          ≃⟨ ≃-≡ (↔⇒≃ W-unfolding) ⟩□
-    (sup x f ≡ sup y g)                          □
-    where
-    lemma : (p : x ≡ y) →
-            (∀ i → f i ≡ g (subst B p i)) ≃
-            (subst (λ x → B x → W A B) p f ≡ g)
-    lemma p = elim
-      (λ {x y} p → (f : B x → W A B) (g : B y → W A B) →
-                   (∀ i → f i ≡ g (subst B p i)) ≃
-                   (subst (λ x → B x → W A B) p f ≡ g))
-      (λ x f g →
-         (∀ i → f i ≡ g (subst B (refl x) i))        ≃⟨ subst (λ h → (∀ i → f i ≡ g (h i)) ≃ (∀ i → f i ≡ g i))
-                                                              (sym (lower-extensionality₂ a ext (subst-refl B)))
-                                                              id ⟩
-         (∀ i → f i ≡ g i)                           ≃⟨ extensionality-isomorphism ext ⟩
-         (f ≡ g)                                     ≃⟨ subst (λ h → (f ≡ g) ≃ (h ≡ g))
-                                                              (sym $ subst-refl (λ x' → B x' → W A B) f)
-                                                              id ⟩□
-         (subst (λ x → B x → W A B) (refl x) f ≡ g)  □)
-      p f g
+-- As a consequence of extensionality-isomorphism and ≃-≡ we get a
+-- strengthening of W-≡,≡↠≡.
+
+W-≡,≡≃≡ : ∀ {a b} {A : Set a} {B : A → Set b} →
+          (∀ {x} {C : B x → Set (a ⊔ b)} → Extensionality′ (B x) C) →
+          ∀ {x y} {f : B x → W A B} {g : B y → W A B} →
+          (∃ λ (p : x ≡ y) → ∀ i → f i ≡ g (subst B p i)) ≃
+          (sup x f ≡ sup y g)
+W-≡,≡≃≡ {a} {A = A} {B} ext {x} {y} {f} {g} =
+  (∃ λ p → ∀ i → f i ≡ g (subst B p i))        ≃⟨ Σ-preserves id lemma ⟩
+  (∃ λ p → subst (λ x → B x → W A B) p f ≡ g)  ≃⟨ ↔⇒≃ Σ-≡,≡↔≡ ⟩
+  ((x , f) ≡ (y , g))                          ≃⟨ ≃-≡ (↔⇒≃ W-unfolding) ⟩□
+  (sup x f ≡ sup y g)                          □
+  where
+  lemma : (p : x ≡ y) →
+          (∀ i → f i ≡ g (subst B p i)) ≃
+          (subst (λ x → B x → W A B) p f ≡ g)
+  lemma p = elim
+    (λ {x y} p → (f : B x → W A B) (g : B y → W A B) →
+                 (∀ i → f i ≡ g (subst B p i)) ≃
+                 (subst (λ x → B x → W A B) p f ≡ g))
+    (λ x f g →
+       (∀ i → f i ≡ g (subst B (refl x) i))        ≃⟨ subst (λ h → (∀ i → f i ≡ g (h i)) ≃ (∀ i → f i ≡ g i))
+                                                            (sym (lower-extensionality₂ a ext (subst-refl B)))
+                                                            id ⟩
+       (∀ i → f i ≡ g i)                           ≃⟨ extensionality-isomorphism ext ⟩
+       (f ≡ g)                                     ≃⟨ subst (λ h → (f ≡ g) ≃ (h ≡ g))
+                                                            (sym $ subst-refl (λ x' → B x' → W A B) f)
+                                                            id ⟩□
+       (subst (λ x → B x → W A B) (refl x) f ≡ g)  □)
+    p f g
hunk ./Fin.agda 38
-abstract
-
-  Fin-set : ∀ n → Is-set (Fin n)
-  Fin-set zero    = mono₁ 1 ⊥-propositional
-  Fin-set (suc n) = ⊎-closure 0 (mono 0≤2 ⊤-contractible) (Fin-set n)
-    where
-    0≤2 : 0 ≤ 2
-    0≤2 = ≤-step (≤-step ≤-refl)
+Fin-set : ∀ n → Is-set (Fin n)
+Fin-set zero    = mono₁ 1 ⊥-propositional
+Fin-set (suc n) = ⊎-closure 0 (mono 0≤2 ⊤-contractible) (Fin-set n)
+  where
+  0≤2 : 0 ≤ 2
+  0≤2 = ≤-step (≤-step ≤-refl)
hunk ./Fin.agda 78
-abstract
-
-  -- The tails are related.
-
-  cancel-suc-preserves-relatedness :
-    ∀ {a} {A : Set a} (x : A) xs ys
-    (f : Fin (length (x ∷ xs)) ↔ Fin (length (x ∷ ys))) →
-    x ∷ xs And x ∷ ys Are-related-by f →
-    xs And ys Are-related-by cancel-suc f
-  cancel-suc-preserves-relatedness x xs ys f related = helper
-    where
-    open _↔_ f
-
-    helper : xs And ys Are-related-by cancel-suc f
-    helper i with inspect (to (fsuc i)) | related (fsuc i)
-    helper i | fsuc k with-≡ eq₁ | hyp₁ =
-      lookup xs i                    ≡⟨ hyp₁ ⟩
-      lookup (x ∷ ys) (to (fsuc i))  ≡⟨ cong (lookup (x ∷ ys)) eq₁ ⟩
-      lookup (x ∷ ys) (fsuc k)       ≡⟨ refl ⟩∎
-      lookup ys k                    ∎
-    helper i | fzero with-≡ eq₁ | hyp₁
-      with inspect (to (fzero)) | related (fzero)
-    helper i | fzero with-≡ eq₁ | hyp₁ | fsuc j with-≡ eq₂ | hyp₂ =
-      lookup xs i                    ≡⟨ hyp₁ ⟩
-      lookup (x ∷ ys) (to (fsuc i))  ≡⟨ cong (lookup (x ∷ ys)) eq₁ ⟩
-      lookup (x ∷ ys) (fzero)        ≡⟨ refl ⟩
-      x                              ≡⟨ hyp₂ ⟩
-      lookup (x ∷ ys) (to (fzero))   ≡⟨ cong (lookup (x ∷ ys)) eq₂ ⟩
-      lookup (x ∷ ys) (fsuc j)       ≡⟨ refl ⟩∎
-      lookup ys j                    ∎
-    helper i | fzero with-≡ eq₁ | hyp₁ | fzero with-≡ eq₂ | hyp₂ =
-      ⊥-elim $ well-behaved f eq₁ eq₂
+-- The tails are related.
+
+cancel-suc-preserves-relatedness :
+  ∀ {a} {A : Set a} (x : A) xs ys
+  (f : Fin (length (x ∷ xs)) ↔ Fin (length (x ∷ ys))) →
+  x ∷ xs And x ∷ ys Are-related-by f →
+  xs And ys Are-related-by cancel-suc f
+cancel-suc-preserves-relatedness x xs ys f related = helper
+  where
+  open _↔_ f
+
+  helper : xs And ys Are-related-by cancel-suc f
+  helper i with inspect (to (fsuc i)) | related (fsuc i)
+  helper i | fsuc k with-≡ eq₁ | hyp₁ =
+    lookup xs i                    ≡⟨ hyp₁ ⟩
+    lookup (x ∷ ys) (to (fsuc i))  ≡⟨ cong (lookup (x ∷ ys)) eq₁ ⟩
+    lookup (x ∷ ys) (fsuc k)       ≡⟨ refl ⟩∎
+    lookup ys k                    ∎
+  helper i | fzero with-≡ eq₁ | hyp₁
+    with inspect (to (fzero)) | related (fzero)
+  helper i | fzero with-≡ eq₁ | hyp₁ | fsuc j with-≡ eq₂ | hyp₂ =
+    lookup xs i                    ≡⟨ hyp₁ ⟩
+    lookup (x ∷ ys) (to (fsuc i))  ≡⟨ cong (lookup (x ∷ ys)) eq₁ ⟩
+    lookup (x ∷ ys) (fzero)        ≡⟨ refl ⟩
+    x                              ≡⟨ hyp₂ ⟩
+    lookup (x ∷ ys) (to (fzero))   ≡⟨ cong (lookup (x ∷ ys)) eq₂ ⟩
+    lookup (x ∷ ys) (fsuc j)       ≡⟨ refl ⟩∎
+    lookup ys j                    ∎
+  helper i | fzero with-≡ eq₁ | hyp₁ | fzero with-≡ eq₂ | hyp₂ =
+    ⊥-elim $ well-behaved f eq₁ eq₂
hunk ./Fin.agda 118
-  abstract
-    to : ∀ m n → (Fin m ↔ Fin n) → m ≡ n
-    to zero    zero    _       = refl
-    to (suc m) (suc n) 1+m↔1+n = cong suc $ to m n $ cancel-suc 1+m↔1+n
-    to zero    (suc n)   0↔1+n = ⊥-elim $ _↔_.from 0↔1+n fzero
-    to (suc m) zero    1+m↔0   = ⊥-elim $ _↔_.to 1+m↔0 fzero
+  to : ∀ m n → (Fin m ↔ Fin n) → m ≡ n
+  to zero    zero    _       = refl
+  to (suc m) (suc n) 1+m↔1+n = cong suc $ to m n $ cancel-suc 1+m↔1+n
+  to zero    (suc n)   0↔1+n = ⊥-elim $ _↔_.from 0↔1+n fzero
+  to (suc m) zero    1+m↔0   = ⊥-elim $ _↔_.to 1+m↔0 fzero
hunk ./Function-universe.agda 252
-abstract
-
-  -- Some lemmas that can be used to manipulate expressions involving
-  -- ≡⇒↝ and refl/sym/trans.
-
-  ≡⇒↝-refl : ∀ {k a} {A : Set a} →
-             ≡⇒↝ k (refl A) ≡ id
-  ≡⇒↝-refl {k} = elim-refl (λ {A B} _ → A ↝[ k ] B) _
-
-  ≡⇒↝-sym : ∀ k {ℓ} {A B : Set ℓ} {eq : A ≡ B} →
-            to-implication (≡⇒↝ ⌊ k ⌋-sym (sym eq)) ≡
-            to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym eq))
-  ≡⇒↝-sym k {A = A} {eq = eq} = elim¹
-    (λ eq → to-implication (≡⇒↝ ⌊ k ⌋-sym (sym eq)) ≡
-            to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym eq)))
-    (to-implication (≡⇒↝ ⌊ k ⌋-sym (sym (refl A)))      ≡⟨ cong (to-implication ∘ ≡⇒↝ ⌊ k ⌋-sym) sym-refl ⟩
-     to-implication (≡⇒↝ ⌊ k ⌋-sym (refl A))            ≡⟨ cong (to-implication {k = ⌊ k ⌋-sym}) ≡⇒↝-refl ⟩
-     to-implication {k = ⌊ k ⌋-sym} id                  ≡⟨ to-implication-id ⌊ k ⌋-sym ⟩
-     id                                                 ≡⟨ sym $ to-implication-inverse-id k ⟩
-     to-implication (inverse {k = k} id)                ≡⟨ cong (to-implication ∘ inverse {k = k}) $ sym ≡⇒↝-refl ⟩∎
-     to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym (refl A)))  ∎)
-    eq
-
-  ≡⇒↝-trans : ∀ k {ℓ} {A B C : Set ℓ} {A≡B : A ≡ B} {B≡C : B ≡ C} →
-              to-implication (≡⇒↝ k (trans A≡B B≡C)) ≡
-              to-implication (≡⇒↝ k B≡C ∘ ≡⇒↝ k A≡B)
-  ≡⇒↝-trans k {B = B} {A≡B = A≡B} = elim¹
-    (λ B≡C → to-implication (≡⇒↝ k (trans A≡B B≡C)) ≡
-             to-implication (≡⇒↝ k B≡C ∘ ≡⇒↝ k A≡B))
-    (to-implication (≡⇒↝ k (trans A≡B (refl B)))             ≡⟨ cong (to-implication ∘ ≡⇒↝ k) $ trans-reflʳ _ ⟩
-     to-implication (≡⇒↝ k A≡B)                              ≡⟨ sym $ cong (λ f → f ∘ to-implication (≡⇒↝ k A≡B)) $ to-implication-id k ⟩
-     to-implication {k = k} id ∘ to-implication (≡⇒↝ k A≡B)  ≡⟨ sym $ to-implication-∘ k ⟩
-     to-implication (id ∘ ≡⇒↝ k A≡B)                         ≡⟨ sym $ cong (λ f → to-implication (f ∘ ≡⇒↝ k A≡B)) ≡⇒↝-refl ⟩∎
-     to-implication (≡⇒↝ k (refl B) ∘ ≡⇒↝ k A≡B)             ∎)
-    _
-
-  -- One can sometimes "push" ≡⇒↝ through cong.
-  --
-  -- This is a generalisation of a lemma due to Thierry Coquand.
-
-  ≡⇒↝-cong : ∀ {k ℓ p A B} {eq : A ≡ B}
-             (P : Set ℓ → Set p)
-             (P-cong : ∀ {A B} → A ↝[ k ] B → P A ↝[ k ] P B) →
-             P-cong (id {A = A}) ≡ id →
-             ≡⇒↝ _ (cong P eq) ≡ P-cong (≡⇒↝ _ eq)
-  ≡⇒↝-cong {eq = eq} P P-cong P-cong-id = elim¹
-    (λ eq → ≡⇒↝ _ (cong P eq) ≡ P-cong (≡⇒↝ _ eq))
-    (≡⇒↝ _ (cong P (refl _))  ≡⟨ cong (≡⇒↝ _) $ cong-refl P ⟩
-     ≡⇒↝ _ (refl _)           ≡⟨ elim-refl (λ {A B} _ → A ↝[ _ ] B) _ ⟩
-     id                       ≡⟨ sym P-cong-id ⟩
-     P-cong id                ≡⟨ cong P-cong $ sym $
-                                   elim-refl (λ {A B} _ → A ↝[ _ ] B) _ ⟩∎
-     P-cong (≡⇒↝ _ (refl _))  ∎)
-    eq
-
-  -- One can express subst in terms of ≡⇒↝.
-
-  subst-in-terms-of-≡⇒↝ :
-    ∀ k {a p} {A : Set a} {x y} (x≡y : x ≡ y) (P : A → Set p) p →
-    subst P x≡y p ≡ to-implication (≡⇒↝ k (cong P x≡y)) p
-  subst-in-terms-of-≡⇒↝ k x≡y P p = elim¹
-
-    (λ eq → subst P eq p ≡ to-implication (≡⇒↝ k (cong P eq)) p)
-
-    (subst P (refl _) p                          ≡⟨ subst-refl P p ⟩
-     p                                           ≡⟨ sym $ cong (_$ p) (to-implication-id k) ⟩
-     to-implication {k = k} id p                 ≡⟨ sym $ cong (λ f → to-implication {k = k} f p) ≡⇒↝-refl ⟩
-     to-implication (≡⇒↝ k (refl _)) p           ≡⟨ sym $ cong (λ eq → to-implication (≡⇒↝ k eq) p) $ cong-refl P ⟩∎
-     to-implication (≡⇒↝ k (cong P (refl _))) p  ∎)
-
-    x≡y
-
-  subst-in-terms-of-inverse∘≡⇒↝ :
-    ∀ k {a p} {A : Set a} {x y} (x≡y : x ≡ y) (P : A → Set p) p →
-    subst P (sym x≡y) p ≡
-    to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym (cong P x≡y))) p
-  subst-in-terms-of-inverse∘≡⇒↝ k x≡y P p =
-    subst P (sym x≡y) p                                      ≡⟨ subst-in-terms-of-≡⇒↝ ⌊ k ⌋-sym (sym x≡y) P p ⟩
-    to-implication (≡⇒↝ ⌊ k ⌋-sym (cong P (sym x≡y))) p      ≡⟨ cong (λ eq → to-implication (≡⇒↝ ⌊ k ⌋-sym eq) p) (cong-sym P _) ⟩
-    to-implication (≡⇒↝ ⌊ k ⌋-sym (sym $ cong P x≡y)) p      ≡⟨ cong (_$ p) (≡⇒↝-sym k) ⟩∎
-    to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym (cong P x≡y))) p  ∎
-
-  to-implication-≡⇒↝ :
-    ∀ k {ℓ} {A B : Set ℓ} (eq : A ≡ B) →
-    to-implication (≡⇒↝ k eq) ≡ ≡⇒↝ implication eq
-  to-implication-≡⇒↝ k =
-    elim (λ eq → to-implication (≡⇒↝ k eq) ≡ ≡⇒↝ implication eq)
-         (λ A → to-implication (≡⇒↝ k (refl A))  ≡⟨ cong to-implication (≡⇒↝-refl {k = k}) ⟩
-                to-implication {k = k} id        ≡⟨ to-implication-id k ⟩
-                id                               ≡⟨ sym ≡⇒↝-refl ⟩∎
-                ≡⇒↝ implication (refl A)         ∎)
+-- Some lemmas that can be used to manipulate expressions involving
+-- ≡⇒↝ and refl/sym/trans.
+
+≡⇒↝-refl : ∀ {k a} {A : Set a} →
+           ≡⇒↝ k (refl A) ≡ id
+≡⇒↝-refl {k} = elim-refl (λ {A B} _ → A ↝[ k ] B) _
+
+≡⇒↝-sym : ∀ k {ℓ} {A B : Set ℓ} {eq : A ≡ B} →
+          to-implication (≡⇒↝ ⌊ k ⌋-sym (sym eq)) ≡
+          to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym eq))
+≡⇒↝-sym k {A = A} {eq = eq} = elim¹
+  (λ eq → to-implication (≡⇒↝ ⌊ k ⌋-sym (sym eq)) ≡
+          to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym eq)))
+  (to-implication (≡⇒↝ ⌊ k ⌋-sym (sym (refl A)))      ≡⟨ cong (to-implication ∘ ≡⇒↝ ⌊ k ⌋-sym) sym-refl ⟩
+   to-implication (≡⇒↝ ⌊ k ⌋-sym (refl A))            ≡⟨ cong (to-implication {k = ⌊ k ⌋-sym}) ≡⇒↝-refl ⟩
+   to-implication {k = ⌊ k ⌋-sym} id                  ≡⟨ to-implication-id ⌊ k ⌋-sym ⟩
+   id                                                 ≡⟨ sym $ to-implication-inverse-id k ⟩
+   to-implication (inverse {k = k} id)                ≡⟨ cong (to-implication ∘ inverse {k = k}) $ sym ≡⇒↝-refl ⟩∎
+   to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym (refl A)))  ∎)
+  eq
+
+≡⇒↝-trans : ∀ k {ℓ} {A B C : Set ℓ} {A≡B : A ≡ B} {B≡C : B ≡ C} →
+            to-implication (≡⇒↝ k (trans A≡B B≡C)) ≡
+            to-implication (≡⇒↝ k B≡C ∘ ≡⇒↝ k A≡B)
+≡⇒↝-trans k {B = B} {A≡B = A≡B} = elim¹
+  (λ B≡C → to-implication (≡⇒↝ k (trans A≡B B≡C)) ≡
+           to-implication (≡⇒↝ k B≡C ∘ ≡⇒↝ k A≡B))
+  (to-implication (≡⇒↝ k (trans A≡B (refl B)))             ≡⟨ cong (to-implication ∘ ≡⇒↝ k) $ trans-reflʳ _ ⟩
+   to-implication (≡⇒↝ k A≡B)                              ≡⟨ sym $ cong (λ f → f ∘ to-implication (≡⇒↝ k A≡B)) $ to-implication-id k ⟩
+   to-implication {k = k} id ∘ to-implication (≡⇒↝ k A≡B)  ≡⟨ sym $ to-implication-∘ k ⟩
+   to-implication (id ∘ ≡⇒↝ k A≡B)                         ≡⟨ sym $ cong (λ f → to-implication (f ∘ ≡⇒↝ k A≡B)) ≡⇒↝-refl ⟩∎
+   to-implication (≡⇒↝ k (refl B) ∘ ≡⇒↝ k A≡B)             ∎)
+  _
+
+-- One can sometimes "push" ≡⇒↝ through cong.
+--
+-- This is a generalisation of a lemma due to Thierry Coquand.
+
+≡⇒↝-cong : ∀ {k ℓ p A B} {eq : A ≡ B}
+           (P : Set ℓ → Set p)
+           (P-cong : ∀ {A B} → A ↝[ k ] B → P A ↝[ k ] P B) →
+           P-cong (id {A = A}) ≡ id →
+           ≡⇒↝ _ (cong P eq) ≡ P-cong (≡⇒↝ _ eq)
+≡⇒↝-cong {eq = eq} P P-cong P-cong-id = elim¹
+  (λ eq → ≡⇒↝ _ (cong P eq) ≡ P-cong (≡⇒↝ _ eq))
+  (≡⇒↝ _ (cong P (refl _))  ≡⟨ cong (≡⇒↝ _) $ cong-refl P ⟩
+   ≡⇒↝ _ (refl _)           ≡⟨ elim-refl (λ {A B} _ → A ↝[ _ ] B) _ ⟩
+   id                       ≡⟨ sym P-cong-id ⟩
+   P-cong id                ≡⟨ cong P-cong $ sym $
+                                 elim-refl (λ {A B} _ → A ↝[ _ ] B) _ ⟩∎
+   P-cong (≡⇒↝ _ (refl _))  ∎)
+  eq
+
+-- One can express subst in terms of ≡⇒↝.
+
+subst-in-terms-of-≡⇒↝ :
+  ∀ k {a p} {A : Set a} {x y} (x≡y : x ≡ y) (P : A → Set p) p →
+  subst P x≡y p ≡ to-implication (≡⇒↝ k (cong P x≡y)) p
+subst-in-terms-of-≡⇒↝ k x≡y P p = elim¹
+
+  (λ eq → subst P eq p ≡ to-implication (≡⇒↝ k (cong P eq)) p)
+
+  (subst P (refl _) p                          ≡⟨ subst-refl P p ⟩
+   p                                           ≡⟨ sym $ cong (_$ p) (to-implication-id k) ⟩
+   to-implication {k = k} id p                 ≡⟨ sym $ cong (λ f → to-implication {k = k} f p) ≡⇒↝-refl ⟩
+   to-implication (≡⇒↝ k (refl _)) p           ≡⟨ sym $ cong (λ eq → to-implication (≡⇒↝ k eq) p) $ cong-refl P ⟩∎
+   to-implication (≡⇒↝ k (cong P (refl _))) p  ∎)
+
+  x≡y
+
+subst-in-terms-of-inverse∘≡⇒↝ :
+  ∀ k {a p} {A : Set a} {x y} (x≡y : x ≡ y) (P : A → Set p) p →
+  subst P (sym x≡y) p ≡
+  to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym (cong P x≡y))) p
+subst-in-terms-of-inverse∘≡⇒↝ k x≡y P p =
+  subst P (sym x≡y) p                                      ≡⟨ subst-in-terms-of-≡⇒↝ ⌊ k ⌋-sym (sym x≡y) P p ⟩
+  to-implication (≡⇒↝ ⌊ k ⌋-sym (cong P (sym x≡y))) p      ≡⟨ cong (λ eq → to-implication (≡⇒↝ ⌊ k ⌋-sym eq) p) (cong-sym P _) ⟩
+  to-implication (≡⇒↝ ⌊ k ⌋-sym (sym $ cong P x≡y)) p      ≡⟨ cong (_$ p) (≡⇒↝-sym k) ⟩∎
+  to-implication (inverse (≡⇒↝ ⌊ k ⌋-sym (cong P x≡y))) p  ∎
+
+to-implication-≡⇒↝ :
+  ∀ k {ℓ} {A B : Set ℓ} (eq : A ≡ B) →
+  to-implication (≡⇒↝ k eq) ≡ ≡⇒↝ implication eq
+to-implication-≡⇒↝ k =
+  elim (λ eq → to-implication (≡⇒↝ k eq) ≡ ≡⇒↝ implication eq)
+       (λ A → to-implication (≡⇒↝ k (refl A))  ≡⟨ cong to-implication (≡⇒↝-refl {k = k}) ⟩
+              to-implication {k = k} id        ≡⟨ to-implication-id k ⟩
+              id                               ≡⟨ sym ≡⇒↝-refl ⟩∎
+              ≡⇒↝ implication (refl A)         ∎)
hunk ./Function-universe.agda 369
-    abstract
-      injective′ : Injective to′
-      injective′ {x = inj₁ x} {y = inj₁ y} = cong inj₁ ⊚ injective A₁↣A₂ ⊚ ⊎.cancel-inj₁
-      injective′ {x = inj₂ x} {y = inj₂ y} = cong inj₂ ⊚ injective B₁↣B₂ ⊚ ⊎.cancel-inj₂
-      injective′ {x = inj₁ x} {y = inj₂ y} = ⊥-elim ⊚ ⊎.inj₁≢inj₂
-      injective′ {x = inj₂ x} {y = inj₁ y} = ⊥-elim ⊚ ⊎.inj₁≢inj₂ ⊚ sym
+    injective′ : Injective to′
+    injective′ {x = inj₁ x} {y = inj₁ y} = cong inj₁ ⊚ injective A₁↣A₂ ⊚ ⊎.cancel-inj₁
+    injective′ {x = inj₂ x} {y = inj₂ y} = cong inj₂ ⊚ injective B₁↣B₂ ⊚ ⊎.cancel-inj₂
+    injective′ {x = inj₁ x} {y = inj₂ y} = ⊥-elim ⊚ ⊎.inj₁≢inj₂
+    injective′ {x = inj₂ x} {y = inj₁ y} = ⊥-elim ⊚ ⊎.inj₁≢inj₂ ⊚ sym
hunk ./Function-universe.agda 501
-    abstract
-      injective′ : Injective to′
-      injective′ to′-x≡to′-y =
-        cong₂ _,_ (injective A₁↣A₂ (cong proj₁ to′-x≡to′-y))
-                  (injective B₁↣B₂ (cong proj₂ to′-x≡to′-y))
+    injective′ : Injective to′
+    injective′ to′-x≡to′-y =
+      cong₂ _,_ (injective A₁↣A₂ (cong proj₁ to′-x≡to′-y))
+                (injective B₁↣B₂ (cong proj₂ to′-x≡to′-y))
hunk ./Function-universe.agda 886
-    abstract
-      right-inv :
-        ∀ f → _⇔_.to logical-equiv (_⇔_.from logical-equiv f) ≡ f
-      right-inv f = lower-extensionality a c ext λ x →
-        to C↔D (from C↔D (f (to A↔B (from A↔B x))))  ≡⟨ right-inverse-of C↔D _ ⟩
-        f (to A↔B (from A↔B x))                      ≡⟨ cong f $ right-inverse-of A↔B _ ⟩∎
-        f x                                          ∎
-
-      left-inv :
-        ∀ f → _⇔_.from logical-equiv (_⇔_.to logical-equiv f) ≡ f
-      left-inv f = lower-extensionality b d ext λ x →
-        from C↔D (to C↔D (f (from A↔B (to A↔B x))))  ≡⟨ left-inverse-of C↔D _ ⟩
-        f (from A↔B (to A↔B x))                      ≡⟨ cong f $ left-inverse-of A↔B _ ⟩∎
-        f x                                          ∎
+    right-inv :
+      ∀ f → _⇔_.to logical-equiv (_⇔_.from logical-equiv f) ≡ f
+    right-inv f = lower-extensionality a c ext λ x →
+      to C↔D (from C↔D (f (to A↔B (from A↔B x))))  ≡⟨ right-inverse-of C↔D _ ⟩
+      f (to A↔B (from A↔B x))                      ≡⟨ cong f $ right-inverse-of A↔B _ ⟩∎
+      f x                                          ∎
+
+    left-inv :
+      ∀ f → _⇔_.from logical-equiv (_⇔_.to logical-equiv f) ≡ f
+    left-inv f = lower-extensionality b d ext λ x →
+      from C↔D (to C↔D (f (from A↔B (to A↔B x))))  ≡⟨ left-inverse-of C↔D _ ⟩
+      f (from A↔B (to A↔B x))                      ≡⟨ cong f $ left-inverse-of A↔B _ ⟩∎
+      f x                                          ∎
hunk ./Function-universe.agda 916
-    abstract
-      right-inverse-of : ∀ x → to (from x) ≡ x
-      right-inverse-of = _≃_.right-inverse-of A→B≃C→D
-
-      irrelevance : ∀ y (p : to ⁻¹ y) →
-                    (from y , right-inverse-of y) ≡ p
-      irrelevance = _≃_.irrelevance A→B≃C→D
+    right-inverse-of : ∀ x → to (from x) ≡ x
+    right-inverse-of = _≃_.right-inverse-of A→B≃C→D
+
+    irrelevance : ∀ y (p : to ⁻¹ y) →
+                  (from y , right-inverse-of y) ≡ p
+    irrelevance = _≃_.irrelevance A→B≃C→D
hunk ./Function-universe.agda 987
-  abstract
-    to∘from : ∀ x≡y → to (from x≡y) ≡ x≡y
-    to∘from x≡y =
-      to (from x≡y)       ≡⟨ sym $ cong (λ f → f (refl x)) $ to-implication∘from-isomorphism bijection ⌊ k ⌋-iso ⟩
-      trans (refl x) x≡y  ≡⟨ trans-reflˡ _ ⟩∎
-      x≡y                 ∎
+  to∘from : ∀ x≡y → to (from x≡y) ≡ x≡y
+  to∘from x≡y =
+    to (from x≡y)       ≡⟨ sym $ cong (λ f → f (refl x)) $ to-implication∘from-isomorphism bijection ⌊ k ⌋-iso ⟩
+    trans (refl x) x≡y  ≡⟨ trans-reflˡ _ ⟩∎
+    x≡y                 ∎
hunk ./Function-universe.agda 1008
-  abstract
-    from∘to : ∀ f → from (to f) ≡ f
-    from∘to f = ext λ z → Eq.lift-equality ext $ ext λ z≡x →
-      trans z≡x (_≃_.to (f x) (refl x))  ≡⟨ elim (λ {u v} u≡v →
-                                                    (f : ∀ z → (z ≡ v) ≃ (z ≡ y)) →
-                                                    trans u≡v (_≃_.to (f v) (refl v)) ≡
-                                                    _≃_.to (f u) u≡v)
-                                                 (λ _ _ → trans-reflˡ _)
-                                                 z≡x f ⟩∎
-      _≃_.to (f z) z≡x                   ∎
+  from∘to : ∀ f → from (to f) ≡ f
+  from∘to f = ext λ z → Eq.lift-equality ext $ ext λ z≡x →
+    trans z≡x (_≃_.to (f x) (refl x))  ≡⟨ elim (λ {u v} u≡v →
+                                                  (f : ∀ z → (z ≡ v) ≃ (z ≡ y)) →
+                                                  trans u≡v (_≃_.to (f v) (refl v)) ≡
+                                                  _≃_.to (f u) u≡v)
+                                               (λ _ _ → trans-reflˡ _)
+                                               z≡x f ⟩∎
+    _≃_.to (f z) z≡x                   ∎
hunk ./Function-universe.agda 1042
-  abstract
-
-    from∘to : ∀ b → to b x (refl x) ≡ b
-    from∘to b =
-      subst (uncurry B)
-            (proj₂ (other-singleton-contractible x) (x , refl x)) b  ≡⟨ cong (λ eq → subst (uncurry B) eq b) $
-                                                                             other-singleton-contractible-refl x ⟩
-      subst (uncurry B) (refl (x , refl x)) b                        ≡⟨ subst-refl (uncurry B) _ ⟩∎
-      b                                                              ∎
-
-    to∘from : ∀ f → to (f x (refl x)) ≡ f
-    to∘from f = ext λ y → lower-extensionality lzero a ext λ x≡y →
-      elim¹ (λ {y} x≡y →
-               subst (uncurry B)
-                     (proj₂ (other-singleton-contractible x) (y , x≡y))
-                     (f x (refl x)) ≡
-               f y x≡y)
-            (subst (uncurry B)
-                   (proj₂ (other-singleton-contractible x) (x , refl x))
-                   (f x (refl x))                                         ≡⟨ from∘to (f x (refl x)) ⟩∎
-             f x (refl x)                                                 ∎)
-            x≡y
+  from∘to : ∀ b → to b x (refl x) ≡ b
+  from∘to b =
+    subst (uncurry B)
+          (proj₂ (other-singleton-contractible x) (x , refl x)) b  ≡⟨ cong (λ eq → subst (uncurry B) eq b) $
+                                                                           other-singleton-contractible-refl x ⟩
+    subst (uncurry B) (refl (x , refl x)) b                        ≡⟨ subst-refl (uncurry B) _ ⟩∎
+    b                                                              ∎
+
+  to∘from : ∀ f → to (f x (refl x)) ≡ f
+  to∘from f = ext λ y → lower-extensionality lzero a ext λ x≡y →
+    elim¹ (λ {y} x≡y →
+             subst (uncurry B)
+                   (proj₂ (other-singleton-contractible x) (y , x≡y))
+                   (f x (refl x)) ≡
+             f y x≡y)
+          (subst (uncurry B)
+                 (proj₂ (other-singleton-contractible x) (x , refl x))
+                 (f x (refl x))                                         ≡⟨ from∘to (f x (refl x)) ⟩∎
+           f x (refl x)                                                 ∎)
+          x≡y
hunk ./Function-universe.agda 1122
-    abstract
-      injective′ : Injective to′
-      injective′ = cong lift ⊚ injective ⊚ cong lower
+    injective′ : Injective to′
+    injective′ = cong lift ⊚ injective ⊚ cong lower
hunk ./Function-universe.agda 1137
-    abstract
-      right-inverse-of′ : ∀ x →
-                          _⇔_.to logical-equivalence′
-                            (_⇔_.from logical-equivalence′ x) ≡
-                          x
-      right-inverse-of′ = cong lift ⊚ right-inverse-of ⊚ lower
+    right-inverse-of′ : ∀ x →
+                        _⇔_.to logical-equivalence′
+                          (_⇔_.from logical-equivalence′ x) ≡
+                        x
+    right-inverse-of′ = cong lift ⊚ right-inverse-of ⊚ lower
hunk ./Function-universe.agda 1155
-    abstract
-      left-inverse-of′ :
-        ∀ x → _↠_.from surjection′ (_↠_.to surjection′ x) ≡ x
-      left-inverse-of′ = cong lift ⊚ left-inverse-of ⊚ lower
+    left-inverse-of′ :
+      ∀ x → _↠_.from surjection′ (_↠_.to surjection′ x) ≡ x
+    left-inverse-of′ = cong lift ⊚ left-inverse-of ⊚ lower
hunk ./Function-universe.agda 1278
-  abstract
-
-    g∘g : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
-          (f : A ⊎ B ↔ A ⊎ C) →
-          (to-hyp   : Well-behaved (to   f)) →
-          (from-hyp : Well-behaved (from f)) →
-          ∀ b → g (from f) from-hyp (g (to f) to-hyp b) ≡ b
-    g∘g f to-hyp from-hyp b = g∘g′
-      where
-      g∘g′ : g (from f) from-hyp (g (to f) to-hyp b) ≡ b
-      g∘g′ with inspect (to f (inj₂ b))
-      g∘g′ | inj₂ c with-≡ eq₁ with inspect (from f (inj₂ c))
-      g∘g′ | inj₂ c with-≡ eq₁ | inj₂ b′ with-≡ eq₂ = ⊎.cancel-inj₂ (
-                                                        inj₂ b′          ≡⟨ sym eq₂ ⟩
-                                                        from f (inj₂ c)  ≡⟨ to-from f eq₁ ⟩∎
-                                                        inj₂ b           ∎)
-      g∘g′ | inj₂ c with-≡ eq₁ | inj₁ a  with-≡ eq₂ = ⊥-elim $ ⊎.inj₁≢inj₂ (
-                                                        inj₁ a           ≡⟨ sym eq₂ ⟩
-                                                        from f (inj₂ c)  ≡⟨ to-from f eq₁ ⟩∎
-                                                        inj₂ b           ∎)
-      g∘g′ | inj₁ a with-≡ eq₁ with inspect (to f (inj₁ a))
-      g∘g′ | inj₁ a with-≡ eq₁ | inj₁ a′ with-≡ eq₂ = ⊥-elim $ to-hyp eq₁ eq₂
-      g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ with inspect (from f (inj₂ c))
-      g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ | inj₂ b′ with-≡ eq₃ = ⊥-elim $ ⊎.inj₁≢inj₂ (
-                                                                             inj₁ a           ≡⟨ sym $ to-from f eq₂ ⟩
-                                                                             from f (inj₂ c)  ≡⟨ eq₃ ⟩∎
-                                                                             inj₂ b′          ∎)
-      g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ | inj₁ a′ with-≡ eq₃ with inspect (from f (inj₁ a′))
-      g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ | inj₁ a′ with-≡ eq₃ | inj₁ a″ with-≡ eq₄ = ⊥-elim $ from-hyp eq₃ eq₄
-      g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ | inj₁ a′ with-≡ eq₃ | inj₂ b′ with-≡ eq₄ = ⊎.cancel-inj₂ (
-        let lemma =
-              inj₁ a′          ≡⟨ sym eq₃ ⟩
-              from f (inj₂ c)  ≡⟨ to-from f eq₂ ⟩∎
-              inj₁ a           ∎
-        in
-        inj₂ b′           ≡⟨ sym eq₄ ⟩
-        from f (inj₁ a′)  ≡⟨ cong (from f ⊚ inj₁) $ ⊎.cancel-inj₁ lemma ⟩
-        from f (inj₁ a)   ≡⟨ to-from f eq₁ ⟩∎
-        inj₂ b            ∎)
+  g∘g : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
+        (f : A ⊎ B ↔ A ⊎ C) →
+        (to-hyp   : Well-behaved (to   f)) →
+        (from-hyp : Well-behaved (from f)) →
+        ∀ b → g (from f) from-hyp (g (to f) to-hyp b) ≡ b
+  g∘g f to-hyp from-hyp b = g∘g′
+    where
+    g∘g′ : g (from f) from-hyp (g (to f) to-hyp b) ≡ b
+    g∘g′ with inspect (to f (inj₂ b))
+    g∘g′ | inj₂ c with-≡ eq₁ with inspect (from f (inj₂ c))
+    g∘g′ | inj₂ c with-≡ eq₁ | inj₂ b′ with-≡ eq₂ = ⊎.cancel-inj₂ (
+                                                      inj₂ b′          ≡⟨ sym eq₂ ⟩
+                                                      from f (inj₂ c)  ≡⟨ to-from f eq₁ ⟩∎
+                                                      inj₂ b           ∎)
+    g∘g′ | inj₂ c with-≡ eq₁ | inj₁ a  with-≡ eq₂ = ⊥-elim $ ⊎.inj₁≢inj₂ (
+                                                      inj₁ a           ≡⟨ sym eq₂ ⟩
+                                                      from f (inj₂ c)  ≡⟨ to-from f eq₁ ⟩∎
+                                                      inj₂ b           ∎)
+    g∘g′ | inj₁ a with-≡ eq₁ with inspect (to f (inj₁ a))
+    g∘g′ | inj₁ a with-≡ eq₁ | inj₁ a′ with-≡ eq₂ = ⊥-elim $ to-hyp eq₁ eq₂
+    g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ with inspect (from f (inj₂ c))
+    g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ | inj₂ b′ with-≡ eq₃ = ⊥-elim $ ⊎.inj₁≢inj₂ (
+                                                                           inj₁ a           ≡⟨ sym $ to-from f eq₂ ⟩
+                                                                           from f (inj₂ c)  ≡⟨ eq₃ ⟩∎
+                                                                           inj₂ b′          ∎)
+    g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ | inj₁ a′ with-≡ eq₃ with inspect (from f (inj₁ a′))
+    g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ | inj₁ a′ with-≡ eq₃ | inj₁ a″ with-≡ eq₄ = ⊥-elim $ from-hyp eq₃ eq₄
+    g∘g′ | inj₁ a with-≡ eq₁ | inj₂ c  with-≡ eq₂ | inj₁ a′ with-≡ eq₃ | inj₂ b′ with-≡ eq₄ = ⊎.cancel-inj₂ (
+      let lemma =
+            inj₁ a′          ≡⟨ sym eq₃ ⟩
+            from f (inj₂ c)  ≡⟨ to-from f eq₂ ⟩∎
+            inj₁ a           ∎
+      in
+      inj₂ b′           ≡⟨ sym eq₄ ⟩
+      from f (inj₁ a′)  ≡⟨ cong (from f ⊚ inj₁) $ ⊎.cancel-inj₁ lemma ⟩
+      from f (inj₁ a)   ≡⟨ to-from f eq₁ ⟩∎
+      inj₂ b            ∎)
hunk ./Groupoid.agda 41
-  abstract
-
-    -- The identity is an identity for the inverse operator as well.
-
-    identity : ∀ {x} → id {x = x} ⁻¹ ≡ id
-    identity =
-      id ⁻¹       ≡⟨ sym $ right-identity (id ⁻¹) ⟩
-      id ⁻¹ ∘ id  ≡⟨ left-inverse id ⟩∎
-      id          ∎
-
-    -- The inverse operator is involutive.
-
-    involutive : ∀ {x y} (p : x ∼ y) → p ⁻¹ ⁻¹ ≡ p
-    involutive p =
-      p ⁻¹ ⁻¹               ≡⟨ sym $ right-identity (p ⁻¹ ⁻¹) ⟩
-      p ⁻¹ ⁻¹ ∘ id          ≡⟨ sym $ cong (_∘_ (p ⁻¹ ⁻¹)) (left-inverse p) ⟩
-      p ⁻¹ ⁻¹ ∘ (p ⁻¹ ∘ p)  ≡⟨ assoc _ _ _ ⟩
-      (p ⁻¹ ⁻¹ ∘ p ⁻¹) ∘ p  ≡⟨ cong (λ q → q ∘ p) (left-inverse (p ⁻¹)) ⟩
-      id ∘ p                ≡⟨ left-identity p ⟩∎
-      p                     ∎
+  -- The identity is an identity for the inverse operator as well.
+
+  identity : ∀ {x} → id {x = x} ⁻¹ ≡ id
+  identity =
+    id ⁻¹       ≡⟨ sym $ right-identity (id ⁻¹) ⟩
+    id ⁻¹ ∘ id  ≡⟨ left-inverse id ⟩∎
+    id          ∎
+
+  -- The inverse operator is involutive.
+
+  involutive : ∀ {x y} (p : x ∼ y) → p ⁻¹ ⁻¹ ≡ p
+  involutive p =
+    p ⁻¹ ⁻¹               ≡⟨ sym $ right-identity (p ⁻¹ ⁻¹) ⟩
+    p ⁻¹ ⁻¹ ∘ id          ≡⟨ sym $ cong (_∘_ (p ⁻¹ ⁻¹)) (left-inverse p) ⟩
+    p ⁻¹ ⁻¹ ∘ (p ⁻¹ ∘ p)  ≡⟨ assoc _ _ _ ⟩
+    (p ⁻¹ ⁻¹ ∘ p ⁻¹) ∘ p  ≡⟨ cong (λ q → q ∘ p) (left-inverse (p ⁻¹)) ⟩
+    id ∘ p                ≡⟨ left-identity p ⟩∎
+    p                     ∎
hunk ./H-level.agda 55
-abstract
-
-  -- H-level is upwards closed in its first argument.
-
-  mono₁ : ∀ {a} {A : Set a} n → H-level n A → H-level (1 + n) A
-  mono₁         (suc n) h x y = mono₁ n (h x y)
-  mono₁ {A = A} zero    h x y = (trivial x y , irr)
-    where
-    trivial : (x y : A) → x ≡ y
-    trivial x y =
-      x        ≡⟨ sym $ proj₂ h x ⟩
-      proj₁ h  ≡⟨ proj₂ h y ⟩∎
-      y        ∎
-
-    irr : ∀ {x y} (x≡y : x ≡ y) → trivial x y ≡ x≡y
-    irr = elim (λ {x y} x≡y → trivial x y ≡ x≡y)
-               (λ x → trans-symˡ (proj₂ h x))
-
-  mono : ∀ {a m n} {A : Set a} → m ≤ n → H-level m A → H-level n A
-  mono ≤-refl               = id
-  mono (≤-step {n = n} m≤n) = λ h → mono₁ n (mono m≤n h)
-
-  -- If something is contractible given the assumption that it is
-  -- inhabited, then it is propositional.
-
-  [inhabited⇒contractible]⇒propositional :
-    ∀ {a} {A : Set a} → (A → Contractible A) → Is-proposition A
-  [inhabited⇒contractible]⇒propositional h x = mono₁ 0 (h x) x
-
-  -- If something has h-level (1 + n) given the assumption that it is
-  -- inhabited, then it has h-level (1 + n).
-
-  [inhabited⇒+]⇒+ :
-    ∀ {a} {A : Set a} n → (A → H-level (1 + n) A) → H-level (1 + n) A
-  [inhabited⇒+]⇒+ n h x = h x x
-
-  -- Being propositional is logically equivalent to having at most one
-  -- element.
-
-  propositional⇔irrelevant :
-    ∀ {a} {A : Set a} → Is-proposition A ⇔ Proof-irrelevant A
-  propositional⇔irrelevant {A} = record
-    { to   = λ h x y → proj₁ (h x y)
-    ; from = λ irr →
-        [inhabited⇒contractible]⇒propositional (λ x → (x , irr x))
-    }
+-- H-level is upwards closed in its first argument.
+
+mono₁ : ∀ {a} {A : Set a} n → H-level n A → H-level (1 + n) A
+mono₁         (suc n) h x y = mono₁ n (h x y)
+mono₁ {A = A} zero    h x y = (trivial x y , irr)
+  where
+  trivial : (x y : A) → x ≡ y
+  trivial x y =
+    x        ≡⟨ sym $ proj₂ h x ⟩
+    proj₁ h  ≡⟨ proj₂ h y ⟩∎
+    y        ∎
+
+  irr : ∀ {x y} (x≡y : x ≡ y) → trivial x y ≡ x≡y
+  irr = elim (λ {x y} x≡y → trivial x y ≡ x≡y)
+             (λ x → trans-symˡ (proj₂ h x))
+
+mono : ∀ {a m n} {A : Set a} → m ≤ n → H-level m A → H-level n A
+mono ≤-refl               = id
+mono (≤-step {n = n} m≤n) = λ h → mono₁ n (mono m≤n h)
+
+-- If something is contractible given the assumption that it is
+-- inhabited, then it is propositional.
+
+[inhabited⇒contractible]⇒propositional :
+  ∀ {a} {A : Set a} → (A → Contractible A) → Is-proposition A
+[inhabited⇒contractible]⇒propositional h x = mono₁ 0 (h x) x
+
+-- If something has h-level (1 + n) given the assumption that it is
+-- inhabited, then it has h-level (1 + n).
+
+[inhabited⇒+]⇒+ :
+  ∀ {a} {A : Set a} n → (A → H-level (1 + n) A) → H-level (1 + n) A
+[inhabited⇒+]⇒+ n h x = h x x
+
+-- Being propositional is logically equivalent to having at most one
+-- element.
+
+propositional⇔irrelevant :
+  ∀ {a} {A : Set a} → Is-proposition A ⇔ Proof-irrelevant A
+propositional⇔irrelevant {A} = record
+  { to   = λ h x y → proj₁ (h x y)
+  ; from = λ irr →
+      [inhabited⇒contractible]⇒propositional (λ x → (x , irr x))
+  }
hunk ./H-level.agda 107
-abstract
-
-  -- Being a set is logically equivalent to having unique identity
-  -- proofs. Note that this means that, assuming that Agda is
-  -- consistent, I cannot prove (inside Agda) that there is any type
-  -- whose minimal h-level is at least three.
-
-  set⇔UIP : ∀ {a} {A : Set a} →
-            Is-set A ⇔ Uniqueness-of-identity-proofs A
-  set⇔UIP {A = A} = record
-    { to   = λ h {x} {y} x≡y x≡y′ → proj₁ (h x y x≡y x≡y′)
-    ; from = λ UIP x y →
-        [inhabited⇒contractible]⇒propositional (λ x≡y → (x≡y , UIP x≡y))
-    }
+-- Being a set is logically equivalent to having unique identity
+-- proofs. Note that this means that, assuming that Agda is
+-- consistent, I cannot prove (inside Agda) that there is any type
+-- whose minimal h-level is at least three.
+
+set⇔UIP : ∀ {a} {A : Set a} →
+          Is-set A ⇔ Uniqueness-of-identity-proofs A
+set⇔UIP {A = A} = record
+  { to   = λ h {x} {y} x≡y x≡y′ → proj₁ (h x y x≡y x≡y′)
+  ; from = λ UIP x y →
+      [inhabited⇒contractible]⇒propositional (λ x≡y → (x≡y , UIP x≡y))
+  }
hunk ./H-level.agda 129
-  abstract
-    irr′ : ∀ y → to x ≡ y
-    irr′ = λ y →
-      to x         ≡⟨ cong to (irr (from y)) ⟩
-      to (from y)  ≡⟨ right-inverse-of y ⟩∎
-      y            ∎
+  irr′ : ∀ y → to x ≡ y
+  irr′ = λ y →
+    to x         ≡⟨ cong to (irr (from y)) ⟩
+    to (from y)  ≡⟨ right-inverse-of y ⟩∎
+    y            ∎
hunk ./H-level/Closure.agda 45
-abstract
-
-  -- The empty type is not contractible.
-
-  ¬-⊥-contractible : ∀ {ℓ} → ¬ Contractible (⊥ {ℓ = ℓ})
-  ¬-⊥-contractible = ⊥-elim ∘ proj₁
-
-  -- The empty type is propositional.
-
-  ⊥-propositional : ∀ {ℓ} → Is-proposition (⊥ {ℓ = ℓ})
-  ⊥-propositional =
-    _⇔_.from propositional⇔irrelevant (λ x → ⊥-elim x)
-
-  -- Thus any uninhabited type is also propositional.
-
-  ⊥↔uninhabited : ∀ {a ℓ} {A : Set a} → ¬ A → ⊥ {ℓ = ℓ} ↔ A
-  ⊥↔uninhabited ¬A = record
-    { surjection = record
-      { logical-equivalence = record
-        { to   = ⊥-elim
-        ; from = ⊥-elim ∘ ¬A
-        }
-      ; right-inverse-of = ⊥-elim ∘ ¬A
+-- The empty type is not contractible.
+
+¬-⊥-contractible : ∀ {ℓ} → ¬ Contractible (⊥ {ℓ = ℓ})
+¬-⊥-contractible = ⊥-elim ∘ proj₁
+
+-- The empty type is propositional.
+
+⊥-propositional : ∀ {ℓ} → Is-proposition (⊥ {ℓ = ℓ})
+⊥-propositional =
+  _⇔_.from propositional⇔irrelevant (λ x → ⊥-elim x)
+
+-- Thus any uninhabited type is also propositional.
+
+⊥↔uninhabited : ∀ {a ℓ} {A : Set a} → ¬ A → ⊥ {ℓ = ℓ} ↔ A
+⊥↔uninhabited ¬A = record
+  { surjection = record
+    { logical-equivalence = record
+      { to   = ⊥-elim
+      ; from = ⊥-elim ∘ ¬A
hunk ./H-level/Closure.agda 65
-    ; left-inverse-of = λ ()
+    ; right-inverse-of = ⊥-elim ∘ ¬A
hunk ./H-level/Closure.agda 67
-
-  uninhabited-propositional : ∀ {a} {A : Set a} →
-                              ¬ A → Is-proposition A
-  uninhabited-propositional ¬A =
-    respects-surjection (_↔_.surjection $ ⊥↔uninhabited {ℓ = # 0} ¬A) 1
-                        ⊥-propositional
+  ; left-inverse-of = λ ()
+  }
+
+uninhabited-propositional : ∀ {a} {A : Set a} →
+                            ¬ A → Is-proposition A
+uninhabited-propositional ¬A =
+  respects-surjection (_↔_.surjection $ ⊥↔uninhabited {ℓ = # 0} ¬A) 1
+                      ⊥-propositional
hunk ./H-level/Closure.agda 79
-abstract
-
-  -- The booleans are not propositional.
-
-  ¬-Bool-propositional : ¬ Is-proposition Bool
-  ¬-Bool-propositional propositional =
-    Bool.true≢false $
-    (_⇔_.to propositional⇔irrelevant propositional) true false
-
-  -- The booleans form a set.
-
-  Bool-set : Is-set Bool
-  Bool-set = DUIP.decidable⇒set Bool._≟_
+-- The booleans are not propositional.
+
+¬-Bool-propositional : ¬ Is-proposition Bool
+¬-Bool-propositional propositional =
+  Bool.true≢false $
+  (_⇔_.to propositional⇔irrelevant propositional) true false
+
+-- The booleans form a set.
+
+Bool-set : Is-set Bool
+Bool-set = DUIP.decidable⇒set Bool._≟_
hunk ./H-level/Closure.agda 94
-abstract
-
-  -- ℕ is not propositional.
-
-  ¬-ℕ-propositional : ¬ Is-proposition ℕ
-  ¬-ℕ-propositional ℕ-prop =
-    ℕ.0≢+ $ _⇔_.to propositional⇔irrelevant ℕ-prop 0 1
-
-  -- ℕ is a set.
-
-  ℕ-set : Is-set ℕ
-  ℕ-set = DUIP.decidable⇒set ℕ._≟_
+-- ℕ is not propositional.
+
+¬-ℕ-propositional : ¬ Is-proposition ℕ
+¬-ℕ-propositional ℕ-prop =
+  ℕ.0≢+ $ _⇔_.to propositional⇔irrelevant ℕ-prop 0 1
+
+-- ℕ is a set.
+
+ℕ-set : Is-set ℕ
+ℕ-set = DUIP.decidable⇒set ℕ._≟_
hunk ./H-level/Closure.agda 135
-abstract
-
-  -- Given (generalised) extensionality one can define an
-  -- extensionality proof which is well-behaved.
-
-  extensionality⇒well-behaved-extensionality :
-    ∀ {a b} {A : Set a} →
-    ({B : A → Set b} → Extensionality′ A B) →
-    {B : A → Set b} → Well-behaved-extensionality A B
-  extensionality⇒well-behaved-extensionality {A = A} ext {B} =
-    (λ {_} → ext′) , λ f →
-      ext′ (refl ∘ f)  ≡⟨ trans-symˡ _ ⟩∎
-      refl f           ∎
-    where
-    ext′ : Extensionality′ A B
-    ext′ = to (from ext)
-      where open _⇔_ Π-closure-contractible⇔extensionality
+-- Given (generalised) extensionality one can define an
+-- extensionality proof which is well-behaved.
+
+extensionality⇒well-behaved-extensionality :
+  ∀ {a b} {A : Set a} →
+  ({B : A → Set b} → Extensionality′ A B) →
+  {B : A → Set b} → Well-behaved-extensionality A B
+extensionality⇒well-behaved-extensionality {A = A} ext {B} =
+  (λ {_} → ext′) , λ f →
+    ext′ (refl ∘ f)  ≡⟨ trans-symˡ _ ⟩∎
+    refl f           ∎
+  where
+  ext′ : Extensionality′ A B
+  ext′ = to (from ext)
+    where open _⇔_ Π-closure-contractible⇔extensionality
hunk ./H-level/Closure.agda 159
-abstract
-
-  -- "Evaluation rule" for ext⁻¹.
-
-  ext⁻¹-refl : ∀ {a b} {A : Set a} {B : A → Set b}
-               (f : (x : A) → B x) {x} →
-               ext⁻¹ (refl f) x ≡ refl (f x)
-  ext⁻¹-refl f {x} = cong-refl (λ h → h x) {x = f}
-
-  -- Given extensionality there is a (split) surjection from
-  -- ∀ x → f x ≡ g x to f ≡ g.
-
-  ext-surj : ∀ {a b} {A : Set a} →
-             ({B : A → Set b} → Extensionality′ A B) →
-             {B : A → Set b} {f g : (x : A) → B x} →
-             (∀ x → f x ≡ g x) ↠ (f ≡ g)
-  ext-surj {b = b} {A} ext {B} = record
-    { logical-equivalence = record
-      { to   = to
-      ; from = ext⁻¹
-      }
-    ; right-inverse-of =
-        elim (λ {f g} f≡g → to (ext⁻¹ f≡g) ≡ f≡g) λ h →
-          proj₁ ext′ (ext⁻¹ (refl h))  ≡⟨ cong (proj₁ ext′) (proj₁ ext′ λ _ →
-                                            ext⁻¹-refl h) ⟩
-          proj₁ ext′ (refl ∘ h)        ≡⟨ proj₂ ext′ h ⟩∎
-          refl h                       ∎
+-- "Evaluation rule" for ext⁻¹.
+
+ext⁻¹-refl : ∀ {a b} {A : Set a} {B : A → Set b}
+             (f : (x : A) → B x) {x} →
+             ext⁻¹ (refl f) x ≡ refl (f x)
+ext⁻¹-refl f {x} = cong-refl (λ h → h x) {x = f}
+
+-- Given extensionality there is a (split) surjection from
+-- ∀ x → f x ≡ g x to f ≡ g.
+
+ext-surj : ∀ {a b} {A : Set a} →
+           ({B : A → Set b} → Extensionality′ A B) →
+           {B : A → Set b} {f g : (x : A) → B x} →
+           (∀ x → f x ≡ g x) ↠ (f ≡ g)
+ext-surj {b = b} {A} ext {B} = record
+  { logical-equivalence = record
+    { to   = to
+    ; from = ext⁻¹
hunk ./H-level/Closure.agda 178
-    where
-    ext′ : {B : A → Set b} → Well-behaved-extensionality A B
-    ext′ = extensionality⇒well-behaved-extensionality ext
-
-    to : {f g : (x : A) → B x} → (∀ x → f x ≡ g x) → f ≡ g
-    to = proj₁ ext′
+  ; right-inverse-of =
+      elim (λ {f g} f≡g → to (ext⁻¹ f≡g) ≡ f≡g) λ h →
+        proj₁ ext′ (ext⁻¹ (refl h))  ≡⟨ cong (proj₁ ext′) (proj₁ ext′ λ _ →
+                                          ext⁻¹-refl h) ⟩
+        proj₁ ext′ (refl ∘ h)        ≡⟨ proj₂ ext′ h ⟩∎
+        refl h                       ∎
+  }
+  where
+  ext′ : {B : A → Set b} → Well-behaved-extensionality A B
+  ext′ = extensionality⇒well-behaved-extensionality ext
+
+  to : {f g : (x : A) → B x} → (∀ x → f x ≡ g x) → f ≡ g
+  to = proj₁ ext′
hunk ./H-level/Closure.agda 217
-abstract
-
-  -- Negated types are propositional, assuming extensionality.
-
-  ¬-propositional :
-    ∀ {a} {A : Set a} →
-    ({B : A → Set} → Extensionality′ A B) →
-    Is-proposition (¬ A)
-  ¬-propositional ext = Π-closure ext 1 (λ _ → ⊥-propositional)
+-- Negated types are propositional, assuming extensionality.
+
+¬-propositional :
+  ∀ {a} {A : Set a} →
+  ({B : A → Set} → Extensionality′ A B) →
+  Is-proposition (¬ A)
+¬-propositional ext = Π-closure ext 1 (λ _ → ⊥-propositional)
hunk ./H-level/Closure.agda 228
-abstract
-
-  -- H-level is closed under Σ.
-
-  Σ-closure : ∀ {a b} {A : Set a} {B : A → Set b} n →
-              H-level n A → (∀ x → H-level n (B x)) → H-level n (Σ A B)
-  Σ-closure {A = A} {B} zero (x , irrA) hB =
-    ((x , proj₁ (hB x)) , λ p →
-       (x       , proj₁ (hB x))          ≡⟨ elim (λ {x y} _ → _≡_ {A = Σ A B} (x , proj₁ (hB x))
-                                                                              (y , proj₁ (hB y)))
-                                                 (λ _ → refl _)
-                                                 (irrA (proj₁ p)) ⟩
-       (proj₁ p , proj₁ (hB (proj₁ p)))  ≡⟨ cong (_,_ (proj₁ p)) (proj₂ (hB (proj₁ p)) (proj₂ p)) ⟩∎
-       p                                 ∎)
-  Σ-closure {B = B} (suc n) hA hB = λ p₁ p₂ →
-    respects-surjection (_↔_.surjection Σ-≡,≡↔≡) n $
-      Σ-closure n (hA (proj₁ p₁) (proj₁ p₂))
-        (λ pr₁p₁≡pr₁p₂ →
-           hB (proj₁ p₂) (subst B pr₁p₁≡pr₁p₂ (proj₂ p₁)) (proj₂ p₂))
-
-  -- In the case of contractibility the codomain only needs to have
-  -- the right h-level (0) for a single index.
-
-  Σ-closure-contractible :
-    ∀ {a b} {A : Set a} {B : A → Set b} →
-    (c : Contractible A) → Contractible (B (proj₁ c)) →
-    Contractible (Σ A B)
-  Σ-closure-contractible {B = B} cA (b , irrB) = Σ-closure 0 cA cB
-    where
-    cB : ∀ a → Contractible (B a)
-    cB a =
-      subst B (proj₂ cA a) b , λ b′ →
-
-      subst B (proj₂ cA a) b                                ≡⟨ cong (subst B (proj₂ cA a)) $
-                                                                irrB (subst B (sym $ proj₂ cA a) b′) ⟩
-      subst B (proj₂ cA a) (subst B (sym $ proj₂ cA a) b′)  ≡⟨ subst-subst-sym _ _ _ ⟩∎
-
-      b′                                                    ∎
-
-  -- H-level is closed under _×_.
-
-  ×-closure : ∀ {a b} {A : Set a} {B : Set b} n →
-              H-level n A → H-level n B → H-level n (A × B)
-  ×-closure n hA hB = Σ-closure n hA (const hB)
-
-  -- If B a is inhabited for all a, and Σ A B has h-level n, then A
-  -- also has h-level n.
-
-  proj₁-closure :
-    ∀ {a b} {A : Set a} {B : A → Set b} →
-    (∀ a → B a) →
-    ∀ n → H-level n (Σ A B) → H-level n A
-  proj₁-closure {A = A} {B} inhabited = respects-surjection surj
-    where
-    surj : Σ A B ↠ A
-    surj = record
-      { logical-equivalence = record
-        { to   = proj₁
-        ; from = λ x → x , inhabited x
-        }
-      ; right-inverse-of = refl
+-- H-level is closed under Σ.
+
+Σ-closure : ∀ {a b} {A : Set a} {B : A → Set b} n →
+            H-level n A → (∀ x → H-level n (B x)) → H-level n (Σ A B)
+Σ-closure {A = A} {B} zero (x , irrA) hB =
+  ((x , proj₁ (hB x)) , λ p →
+     (x       , proj₁ (hB x))          ≡⟨ elim (λ {x y} _ → _≡_ {A = Σ A B} (x , proj₁ (hB x))
+                                                                            (y , proj₁ (hB y)))
+                                               (λ _ → refl _)
+                                               (irrA (proj₁ p)) ⟩
+     (proj₁ p , proj₁ (hB (proj₁ p)))  ≡⟨ cong (_,_ (proj₁ p)) (proj₂ (hB (proj₁ p)) (proj₂ p)) ⟩∎
+     p                                 ∎)
+Σ-closure {B = B} (suc n) hA hB = λ p₁ p₂ →
+  respects-surjection (_↔_.surjection Σ-≡,≡↔≡) n $
+    Σ-closure n (hA (proj₁ p₁) (proj₁ p₂))
+      (λ pr₁p₁≡pr₁p₂ →
+         hB (proj₁ p₂) (subst B pr₁p₁≡pr₁p₂ (proj₂ p₁)) (proj₂ p₂))
+
+-- In the case of contractibility the codomain only needs to have
+-- the right h-level (0) for a single index.
+
+Σ-closure-contractible :
+  ∀ {a b} {A : Set a} {B : A → Set b} →
+  (c : Contractible A) → Contractible (B (proj₁ c)) →
+  Contractible (Σ A B)
+Σ-closure-contractible {B = B} cA (b , irrB) = Σ-closure 0 cA cB
+  where
+  cB : ∀ a → Contractible (B a)
+  cB a =
+    subst B (proj₂ cA a) b , λ b′ →
+
+    subst B (proj₂ cA a) b                                ≡⟨ cong (subst B (proj₂ cA a)) $
+                                                              irrB (subst B (sym $ proj₂ cA a) b′) ⟩
+    subst B (proj₂ cA a) (subst B (sym $ proj₂ cA a) b′)  ≡⟨ subst-subst-sym _ _ _ ⟩∎
+
+    b′                                                    ∎
+
+-- H-level is closed under _×_.
+
+×-closure : ∀ {a b} {A : Set a} {B : Set b} n →
+            H-level n A → H-level n B → H-level n (A × B)
+×-closure n hA hB = Σ-closure n hA (const hB)
+
+-- If B a is inhabited for all a, and Σ A B has h-level n, then A
+-- also has h-level n.
+
+proj₁-closure :
+  ∀ {a b} {A : Set a} {B : A → Set b} →
+  (∀ a → B a) →
+  ∀ n → H-level n (Σ A B) → H-level n A
+proj₁-closure {A = A} {B} inhabited = respects-surjection surj
+  where
+  surj : Σ A B ↠ A
+  surj = record
+    { logical-equivalence = record
+      { to   = proj₁
+      ; from = λ x → x , inhabited x
hunk ./H-level/Closure.agda 286
-
-  -- If A is inhabited and A × B has h-level n, then B also has
-  -- h-level n.
-
-  proj₂-closure :
-    ∀ {a b} {A : Set a} {B : Set b} →
-    A →
-    ∀ n → H-level n (A × B) → H-level n B
-  proj₂-closure {A = A} {B} inhabited = respects-surjection surj
-    where
-    surj : A × B ↠ B
-    surj = record
-      { logical-equivalence = record
-        { to   = proj₂
-        ; from = λ x → inhabited , x
-        }
-      ; right-inverse-of = refl
+    ; right-inverse-of = refl
+    }
+
+-- If A is inhabited and A × B has h-level n, then B also has
+-- h-level n.
+
+proj₂-closure :
+  ∀ {a b} {A : Set a} {B : Set b} →
+  A →
+  ∀ n → H-level n (A × B) → H-level n B
+proj₂-closure {A = A} {B} inhabited = respects-surjection surj
+  where
+  surj : A × B ↠ B
+  surj = record
+    { logical-equivalence = record
+      { to   = proj₂
+      ; from = λ x → inhabited , x
hunk ./H-level/Closure.agda 304
+    ; right-inverse-of = refl
+    }
hunk ./H-level/Closure.agda 310
-abstract
-
-  -- All H-levels are closed under lifting.
-
-  ↑-closure : ∀ {a b} {A : Set a} n → H-level n A → H-level n (↑ b A)
-  ↑-closure =
-    respects-surjection (_↔_.surjection (Bijection.inverse ↑↔))
-
-  -- All H-levels are also closed under removal of lifting.
-
-  ↑⁻¹-closure : ∀ {a b} {A : Set a} n → H-level n (↑ b A) → H-level n A
-  ↑⁻¹-closure = respects-surjection (_↔_.surjection ↑↔)
+-- All H-levels are closed under lifting.
+
+↑-closure : ∀ {a b} {A : Set a} n → H-level n A → H-level n (↑ b A)
+↑-closure =
+  respects-surjection (_↔_.surjection (Bijection.inverse ↑↔))
+
+-- All H-levels are also closed under removal of lifting.
+
+↑⁻¹-closure : ∀ {a b} {A : Set a} n → H-level n (↑ b A) → H-level n A
+↑⁻¹-closure = respects-surjection (_↔_.surjection ↑↔)
hunk ./H-level/Closure.agda 342
-abstract
-
-  -- Equality between elements of a W-type can be proved using a pair
-  -- of equalities (assuming extensionality).
-
-  W-≡,≡↠≡ : ∀ {a b} {A : Set a} {B : A → Set b} →
-            (∀ {x} {C : B x → Set (a ⊔ b)} → Extensionality′ (B x) C) →
-            ∀ {x y} {f : B x → W A B} {g : B y → W A B} →
-            (∃ λ (p : x ≡ y) → ∀ i → f i ≡ g (subst B p i)) ↠
-            (sup x f ≡ sup y g)
-  W-≡,≡↠≡ {a} {A = A} {B} ext {x} {y} {f} {g} =
-    (∃ λ (p : x ≡ y) → ∀ i → f i ≡ g (subst B p i))        ↠⟨ Surjection.∃-cong lemma ⟩
-    (∃ λ (p : x ≡ y) → subst (λ x → B x → W A B) p f ≡ g)  ↠⟨ _↔_.surjection Σ-≡,≡↔≡ ⟩
-    (_≡_ {A = ∃ λ (x : A) → B x → W A B} (x , f) (y , g))  ↠⟨ ↠-≡ (_↔_.surjection (Bijection.inverse (W-unfolding {A = A} {B = B}))) ⟩□
-    (sup x f ≡ sup y g)                                    □
-    where
-    lemma : (p : x ≡ y) →
-            (∀ i → f i ≡ g (subst B p i)) ↠
-            (subst (λ x → B x → W A B) p f ≡ g)
-    lemma p = elim
-      (λ {x y} p → (f : B x → W A B) (g : B y → W A B) →
-                   (∀ i → f i ≡ g (subst B p i)) ↠
-                   (subst (λ x → B x → W A B) p f ≡ g))
-      (λ x f g →
-         (∀ i → f i ≡ g (subst B (refl x) i))        ↠⟨ subst (λ h → (∀ i → f i ≡ g (h i)) ↠ (∀ i → f i ≡ g i))
-                                                              (sym (lower-extensionality₂ a ext (subst-refl B)))
-                                                              Surjection.id ⟩
-         (∀ i → f i ≡ g i)                           ↠⟨ ext-surj ext ⟩
-         (f ≡ g)                                     ↠⟨ subst (λ h → (f ≡ g) ↠ (h ≡ g))
-                                                              (sym $ subst-refl (λ x' → B x' → W A B) f)
-                                                              Surjection.id ⟩□
-         (subst (λ x → B x → W A B) (refl x) f ≡ g)  □)
-      p f g
-
-  -- H-level is not closed under W.
-
-  ¬-W-closure-contractible : ∀ {a b} →
-    ¬ (∀ {A : Set a} {B : A → Set b} →
-       Contractible A → (∀ x → Contractible (B x)) →
-       Contractible (W A B))
-  ¬-W-closure-contractible closure =
-    inhabited⇒W-empty (const (lift tt)) $
-    proj₁ $
-    closure (↑-closure 0 ⊤-contractible)
-            (const (↑-closure 0 ⊤-contractible))
-
-  ¬-W-closure : ∀ {a b} →
-    ¬ (∀ {A : Set a} {B : A → Set b} n →
-       H-level n A → (∀ x → H-level n (B x)) → H-level n (W A B))
-  ¬-W-closure closure = ¬-W-closure-contractible (closure 0)
-
-  -- However, all positive h-levels are closed under W, assuming that
-  -- equality is sufficiently extensional.
-
-  W-closure :
-    ∀ {a b} {A : Set a} {B : A → Set b} →
-    (∀ {x} {C : B x → Set (a ⊔ b)} → Extensionality′ (B x) C) →
-    ∀ n → H-level (1 + n) A → H-level (1 + n) (W A B)
-  W-closure {A = A} {B} ext n h = closure
-    where
-    closure : (x y : W A B) → H-level n (x ≡ y)
-    closure (sup x f) (sup y g) =
-      respects-surjection (W-≡,≡↠≡ ext) n $
-        Σ-closure n (h x y) (λ _ →
-          Π-closure ext n (λ i → closure (f _) (g _)))
+-- Equality between elements of a W-type can be proved using a pair
+-- of equalities (assuming extensionality).
+
+W-≡,≡↠≡ : ∀ {a b} {A : Set a} {B : A → Set b} →
+          (∀ {x} {C : B x → Set (a ⊔ b)} → Extensionality′ (B x) C) →
+          ∀ {x y} {f : B x → W A B} {g : B y → W A B} →
+          (∃ λ (p : x ≡ y) → ∀ i → f i ≡ g (subst B p i)) ↠
+          (sup x f ≡ sup y g)
+W-≡,≡↠≡ {a} {A = A} {B} ext {x} {y} {f} {g} =
+  (∃ λ (p : x ≡ y) → ∀ i → f i ≡ g (subst B p i))        ↠⟨ Surjection.∃-cong lemma ⟩
+  (∃ λ (p : x ≡ y) → subst (λ x → B x → W A B) p f ≡ g)  ↠⟨ _↔_.surjection Σ-≡,≡↔≡ ⟩
+  (_≡_ {A = ∃ λ (x : A) → B x → W A B} (x , f) (y , g))  ↠⟨ ↠-≡ (_↔_.surjection (Bijection.inverse (W-unfolding {A = A} {B = B}))) ⟩□
+  (sup x f ≡ sup y g)                                    □
+  where
+  lemma : (p : x ≡ y) →
+          (∀ i → f i ≡ g (subst B p i)) ↠
+          (subst (λ x → B x → W A B) p f ≡ g)
+  lemma p = elim
+    (λ {x y} p → (f : B x → W A B) (g : B y → W A B) →
+                 (∀ i → f i ≡ g (subst B p i)) ↠
+                 (subst (λ x → B x → W A B) p f ≡ g))
+    (λ x f g →
+       (∀ i → f i ≡ g (subst B (refl x) i))        ↠⟨ subst (λ h → (∀ i → f i ≡ g (h i)) ↠ (∀ i → f i ≡ g i))
+                                                            (sym (lower-extensionality₂ a ext (subst-refl B)))
+                                                            Surjection.id ⟩
+       (∀ i → f i ≡ g i)                           ↠⟨ ext-surj ext ⟩
+       (f ≡ g)                                     ↠⟨ subst (λ h → (f ≡ g) ↠ (h ≡ g))
+                                                            (sym $ subst-refl (λ x' → B x' → W A B) f)
+                                                            Surjection.id ⟩□
+       (subst (λ x → B x → W A B) (refl x) f ≡ g)  □)
+    p f g
+
+-- H-level is not closed under W.
+
+¬-W-closure-contractible : ∀ {a b} →
+  ¬ (∀ {A : Set a} {B : A → Set b} →
+     Contractible A → (∀ x → Contractible (B x)) →
+     Contractible (W A B))
+¬-W-closure-contractible closure =
+  inhabited⇒W-empty (const (lift tt)) $
+  proj₁ $
+  closure (↑-closure 0 ⊤-contractible)
+          (const (↑-closure 0 ⊤-contractible))
+
+¬-W-closure : ∀ {a b} →
+  ¬ (∀ {A : Set a} {B : A → Set b} n →
+     H-level n A → (∀ x → H-level n (B x)) → H-level n (W A B))
+¬-W-closure closure = ¬-W-closure-contractible (closure 0)
+
+-- However, all positive h-levels are closed under W, assuming that
+-- equality is sufficiently extensional.
+
+W-closure :
+  ∀ {a b} {A : Set a} {B : A → Set b} →
+  (∀ {x} {C : B x → Set (a ⊔ b)} → Extensionality′ (B x) C) →
+  ∀ n → H-level (1 + n) A → H-level (1 + n) (W A B)
+W-closure {A = A} {B} ext n h = closure
+  where
+  closure : (x y : W A B) → H-level n (x ≡ y)
+  closure (sup x f) (sup y g) =
+    respects-surjection (W-≡,≡↠≡ ext) n $
+      Σ-closure n (h x y) (λ _ →
+        Π-closure ext n (λ i → closure (f _) (g _)))
hunk ./H-level/Closure.agda 409
-abstract
-
-  -- Contractible is /not/ a comonad in the category of types and
-  -- functions, because map cannot be defined, but we can at least
-  -- define the following functions.
-
-  counit : ∀ {a} {A : Set a} → Contractible A → A
-  counit = proj₁
-
-  cojoin : ∀ {a} {A : Set a} →
-           ({B : A → Set a} → Extensionality′ A B) →
-           Contractible A → Contractible (Contractible A)
-  cojoin {A = A} ext contr = contr₃
-    where
-    x : A
-    x = proj₁ contr
-
-    contr₁ : Contractible (∀ y → x ≡ y)
-    contr₁ = Π-closure ext 0 (mono₁ 0 contr x)
-
-    contr₂ : (x : A) → Contractible (∀ y → x ≡ y)
-    contr₂ x =
-      subst (λ x → Contractible (∀ y → x ≡ y)) (proj₂ contr x) contr₁
-
-    contr₃ : Contractible (∃ λ (x : A) → ∀ y → x ≡ y)
-    contr₃ = Σ-closure 0 contr contr₂
-
-  -- Contractible is not necessarily contractible.
-
-  ¬-Contractible-contractible :
-    ¬ ({A : Set} → Contractible (Contractible A))
-  ¬-Contractible-contractible contr = proj₁ $ proj₁ $ contr {A = ⊥}
-
-  -- Contractible is propositional (assuming extensionality).
-
-  Contractible-propositional :
-    ∀ {a} {A : Set a} →
-    ({B : A → Set a} → Extensionality′ A B) →
-    Is-proposition (Contractible A)
-  Contractible-propositional ext =
-    [inhabited⇒contractible]⇒propositional (cojoin ext)
-
-  -- All h-levels are propositional (assuming extensionality).
-
-  H-level-propositional :
-    ∀ {a} → Extensionality a a →
-    ∀ {A : Set a} n → Is-proposition (H-level n A)
-  H-level-propositional     ext zero    = Contractible-propositional ext
-  H-level-propositional {A} ext (suc n) =
-    Π-closure ext 1 λ x →
-    Π-closure ext 1 λ y →
-    H-level-propositional ext {A = x ≡ y} n
+-- Contractible is /not/ a comonad in the category of types and
+-- functions, because map cannot be defined, but we can at least
+-- define the following functions.
+
+counit : ∀ {a} {A : Set a} → Contractible A → A
+counit = proj₁
+
+cojoin : ∀ {a} {A : Set a} →
+         ({B : A → Set a} → Extensionality′ A B) →
+         Contractible A → Contractible (Contractible A)
+cojoin {A = A} ext contr = contr₃
+  where
+  x : A
+  x = proj₁ contr
+
+  contr₁ : Contractible (∀ y → x ≡ y)
+  contr₁ = Π-closure ext 0 (mono₁ 0 contr x)
+
+  contr₂ : (x : A) → Contractible (∀ y → x ≡ y)
+  contr₂ x =
+    subst (λ x → Contractible (∀ y → x ≡ y)) (proj₂ contr x) contr₁
+
+  contr₃ : Contractible (∃ λ (x : A) → ∀ y → x ≡ y)
+  contr₃ = Σ-closure 0 contr contr₂
+
+-- Contractible is not necessarily contractible.
+
+¬-Contractible-contractible :
+  ¬ ({A : Set} → Contractible (Contractible A))
+¬-Contractible-contractible contr = proj₁ $ proj₁ $ contr {A = ⊥}
+
+-- Contractible is propositional (assuming extensionality).
+
+Contractible-propositional :
+  ∀ {a} {A : Set a} →
+  ({B : A → Set a} → Extensionality′ A B) →
+  Is-proposition (Contractible A)
+Contractible-propositional ext =
+  [inhabited⇒contractible]⇒propositional (cojoin ext)
+
+-- All h-levels are propositional (assuming extensionality).
+
+H-level-propositional :
+  ∀ {a} → Extensionality a a →
+  ∀ {A : Set a} n → Is-proposition (H-level n A)
+H-level-propositional     ext zero    = Contractible-propositional ext
+H-level-propositional {A} ext (suc n) =
+  Π-closure ext 1 λ x →
+  Π-closure ext 1 λ y →
+  H-level-propositional ext {A = x ≡ y} n
hunk ./H-level/Closure.agda 463
-abstract
-
-  -- Binary sums can be expressed using Σ and Bool (with large
-  -- elimination).
-
-  sum-as-pair : ∀ {a b} {A : Set a} {B : Set b} →
-                (A ⊎ B) ↔ (∃ λ x → if x then ↑ b A else ↑ a B)
-  sum-as-pair {a} {b} {A} {B} = record
-    { surjection = record
-      { logical-equivalence = record
-        { to   = to
-        ; from = from
-        }
-      ; right-inverse-of = to∘from
+-- Binary sums can be expressed using Σ and Bool (with large
+-- elimination).
+
+sum-as-pair : ∀ {a b} {A : Set a} {B : Set b} →
+              (A ⊎ B) ↔ (∃ λ x → if x then ↑ b A else ↑ a B)
+sum-as-pair {a} {b} {A} {B} = record
+  { surjection = record
+    { logical-equivalence = record
+      { to   = to
+      ; from = from
hunk ./H-level/Closure.agda 474
-    ; left-inverse-of = [ refl ∘ inj₁ {B = B} , refl ∘ inj₂ {A = A} ]
+    ; right-inverse-of = to∘from
hunk ./H-level/Closure.agda 476
+  ; left-inverse-of = [ refl ∘ inj₁ {B = B} , refl ∘ inj₂ {A = A} ]
+  }
+  where
+  to : A ⊎ B → (∃ λ x → if x then ↑ b A else ↑ a B)
+  to = [ _,_ true ∘ lift , _,_ false ∘ lift ]
+
+  from : (∃ λ x → if x then ↑ b A else ↑ a B) → A ⊎ B
+  from (true  , x) = inj₁ $ lower x
+  from (false , y) = inj₂ $ lower y
+
+  to∘from : (y : ∃ λ x → if x then ↑ b A else ↑ a B) →
+            to (from y) ≡ y
+  to∘from (true  , x) = refl _
+  to∘from (false , y) = refl _
+
+-- H-level is not closed under _⊎_.
+
+¬-⊎-propositional : ∀ {a b} {A : Set a} {B : Set b} →
+                    A → B → ¬ Is-proposition (A ⊎ B)
+¬-⊎-propositional {A = A} {B} x y hA⊎B =
+  ⊎.inj₁≢inj₂ {A = A} {B = B} $ proj₁ $ hA⊎B (inj₁ x) (inj₂ y)
+
+¬-⊎-closure : ∀ {a b} →
+  ¬ (∀ {A : Set a} {B : Set b} n →
+     H-level n A → H-level n B → H-level n (A ⊎ B))
+¬-⊎-closure ⊎-closure =
+  ¬-⊎-propositional (lift tt) (lift tt) $
+  mono₁ 0 $
+  ⊎-closure 0 (↑-closure 0 ⊤-contractible)
+              (↑-closure 0 ⊤-contractible)
+
+-- However, all levels greater than or equal to 2 are closed under
+-- _⊎_.
+
+⊎-closure :
+  ∀ {a b} {A : Set a} {B : Set b} n →
+  H-level (2 + n) A → H-level (2 + n) B → H-level (2 + n) (A ⊎ B)
+⊎-closure {a} {b} {A} {B} n hA hB =
+  respects-surjection
+    (_↔_.surjection $ Bijection.inverse sum-as-pair)
+    (2 + n)
+    (Σ-closure (2 + n) Bool-2+n if-2+n)
+  where
+  Bool-2+n : H-level (2 + n) Bool
+  Bool-2+n = mono (m≤m+n 2 n) Bool-set
+
+  if-2+n : ∀ x → H-level (2 + n) (if x then ↑ b A else ↑ a B)
+  if-2+n true  = respects-surjection
+                   (_↔_.surjection $ Bijection.inverse ↑↔)
+                   (2 + n) hA
+  if-2+n false = respects-surjection
+                   (_↔_.surjection $ Bijection.inverse ↑↔)
+                   (2 + n) hB
+
+-- Furthermore Is-proposition is closed under Dec (assuming
+-- extensionality).
+
+Dec-closure-propositional :
+  ∀ {a} {A : Set a} →
+  ({B : A → Set} → Extensionality′ A B) →
+  Is-proposition A → Is-proposition (Dec A)
+Dec-closure-propositional {A = A} ext p =
+  _⇔_.from propositional⇔irrelevant irrelevant
+  where
+  irrelevant : Proof-irrelevant (Dec A)
+  irrelevant (yes  a) (yes  a′) = cong yes $ proj₁ $ p a a′
+  irrelevant (yes  a) (no  ¬a)  = ⊥-elim (¬a a)
+  irrelevant (no  ¬a) (yes  a)  = ⊥-elim (¬a a)
+  irrelevant (no  ¬a) (no  ¬a′) =
+    cong no $ proj₁ $ ¬-propositional ext ¬a ¬a′
+
+-- Is-proposition is also closed under _Xor_ (assuming
+-- extensionality).
+
+Xor-closure-propositional :
+  ∀ {a b} {A : Set a} {B : Set b} →
+  Extensionality (a ⊔ b) (# 0) →
+  Is-proposition A → Is-proposition B →
+  Is-proposition (A Xor B)
+Xor-closure-propositional {ℓa} {ℓb} {A} {B} ext pA pB =
+  _⇔_.from propositional⇔irrelevant irr
+  where
+  irr : (x y : A Xor B) → x ≡ y
+  irr (inj₁ (a , ¬b)) (inj₂ (¬a  , b))   = ⊥-elim (¬a a)
+  irr (inj₂ (¬a , b)) (inj₁ (a   , ¬b))  = ⊥-elim (¬b b)
+  irr (inj₁ (a , ¬b)) (inj₁ (a′  , ¬b′)) =
+    cong₂ (λ x y → inj₁ (x , y))
+      (_⇔_.to propositional⇔irrelevant pA a a′)
+      (lower-extensionality ℓa _ ext λ b → ⊥-elim (¬b b))
+  irr (inj₂ (¬a , b)) (inj₂ (¬a′ , b′)) =
+    cong₂ (λ x y → inj₂ (x , y))
+      (lower-extensionality ℓb _ ext λ a → ⊥-elim (¬a a))
+      (_⇔_.to propositional⇔irrelevant pB b b′)
+
+-- However, H-level is not closed under _Xor_.
+
+¬-Xor-closure-contractible : ∀ {a b} →
+  ¬ ({A : Set a} {B : Set b} →
+     Contractible A → Contractible B → Contractible (A Xor B))
+¬-Xor-closure-contractible closure
+  with proj₁ $ closure (↑-closure 0 ⊤-contractible)
+                       (↑-closure 0 ⊤-contractible)
+... | inj₁ (_ , ¬⊤) = ¬⊤ _
+... | inj₂ (¬⊤ , _) = ¬⊤ _
+
+-- Alternative definition of ⊎-closure (for Set₀).
+
+module Alternative-proof where
+
+  -- Is-set is closed under _⊎_, by an argument similar to the one
+  -- Hedberg used to prove that decidable equality implies
+  -- uniqueness of identity proofs.
+
+  ⊎-closure-set : {A B : Set} →
+                  Is-set A → Is-set B → Is-set (A ⊎ B)
+  ⊎-closure-set {A} {B} A-set B-set =
+    _⇔_.from set⇔UIP (DUIP.constant⇒UIP c)
hunk ./H-level/Closure.agda 594
-    to : A ⊎ B → (∃ λ x → if x then ↑ b A else ↑ a B)
-    to = [ _,_ true ∘ lift , _,_ false ∘ lift ]
-
-    from : (∃ λ x → if x then ↑ b A else ↑ a B) → A ⊎ B
-    from (true  , x) = inj₁ $ lower x
-    from (false , y) = inj₂ $ lower y
-
-    to∘from : (y : ∃ λ x → if x then ↑ b A else ↑ a B) →
-              to (from y) ≡ y
-    to∘from (true  , x) = refl _
-    to∘from (false , y) = refl _
-
-  -- H-level is not closed under _⊎_.
-
-  ¬-⊎-propositional : ∀ {a b} {A : Set a} {B : Set b} →
-                      A → B → ¬ Is-proposition (A ⊎ B)
-  ¬-⊎-propositional {A = A} {B} x y hA⊎B =
-    ⊎.inj₁≢inj₂ {A = A} {B = B} $ proj₁ $ hA⊎B (inj₁ x) (inj₂ y)
-
-  ¬-⊎-closure : ∀ {a b} →
-    ¬ (∀ {A : Set a} {B : Set b} n →
-       H-level n A → H-level n B → H-level n (A ⊎ B))
-  ¬-⊎-closure ⊎-closure =
-    ¬-⊎-propositional (lift tt) (lift tt) $
-    mono₁ 0 $
-    ⊎-closure 0 (↑-closure 0 ⊤-contractible)
-                (↑-closure 0 ⊤-contractible)
-
-  -- However, all levels greater than or equal to 2 are closed under
-  -- _⊎_.
-
-  ⊎-closure :
-    ∀ {a b} {A : Set a} {B : Set b} n →
+    c : (x y : A ⊎ B) → ∃ λ (f : x ≡ y → x ≡ y) → DUIP.Constant f
+    c (inj₁ x) (inj₁ y) = (cong inj₁ ∘ ⊎.cancel-inj₁ , λ p q → cong (cong inj₁) $ proj₁ $ A-set x y (⊎.cancel-inj₁ p) (⊎.cancel-inj₁ q))
+    c (inj₂ x) (inj₂ y) = (cong inj₂ ∘ ⊎.cancel-inj₂ , λ p q → cong (cong inj₂) $ proj₁ $ B-set x y (⊎.cancel-inj₂ p) (⊎.cancel-inj₂ q))
+    c (inj₁ x) (inj₂ y) = (⊥-elim ∘ ⊎.inj₁≢inj₂       , λ _ → ⊥-elim ∘ ⊎.inj₁≢inj₂)
+    c (inj₂ x) (inj₁ y) = (⊥-elim ∘ ⊎.inj₁≢inj₂ ∘ sym , λ _ → ⊥-elim ∘ ⊎.inj₁≢inj₂ ∘ sym)
+
+  -- H-level is closed under _⊎_ for other levels greater than or equal
+  -- to 2 too.
+
+  ⊎-closure′ :
+    ∀ {A B : Set} n →
hunk ./H-level/Closure.agda 606
-  ⊎-closure {a} {b} {A} {B} n hA hB =
-    respects-surjection
-      (_↔_.surjection $ Bijection.inverse sum-as-pair)
-      (2 + n)
-      (Σ-closure (2 + n) Bool-2+n if-2+n)
+  ⊎-closure′         zero    = ⊎-closure-set
+  ⊎-closure′ {A} {B} (suc n) = clos
hunk ./H-level/Closure.agda 609
-    Bool-2+n : H-level (2 + n) Bool
-    Bool-2+n = mono (m≤m+n 2 n) Bool-set
-
-    if-2+n : ∀ x → H-level (2 + n) (if x then ↑ b A else ↑ a B)
-    if-2+n true  = respects-surjection
-                     (_↔_.surjection $ Bijection.inverse ↑↔)
-                     (2 + n) hA
-    if-2+n false = respects-surjection
-                     (_↔_.surjection $ Bijection.inverse ↑↔)
-                     (2 + n) hB
-
-  -- Furthermore Is-proposition is closed under Dec (assuming
-  -- extensionality).
-
-  Dec-closure-propositional :
-    ∀ {a} {A : Set a} →
-    ({B : A → Set} → Extensionality′ A B) →
-    Is-proposition A → Is-proposition (Dec A)
-  Dec-closure-propositional {A = A} ext p =
-    _⇔_.from propositional⇔irrelevant irrelevant
-    where
-    irrelevant : Proof-irrelevant (Dec A)
-    irrelevant (yes  a) (yes  a′) = cong yes $ proj₁ $ p a a′
-    irrelevant (yes  a) (no  ¬a)  = ⊥-elim (¬a a)
-    irrelevant (no  ¬a) (yes  a)  = ⊥-elim (¬a a)
-    irrelevant (no  ¬a) (no  ¬a′) =
-      cong no $ proj₁ $ ¬-propositional ext ¬a ¬a′
-
-  -- Is-proposition is also closed under _Xor_ (assuming
-  -- extensionality).
-
-  Xor-closure-propositional :
-    ∀ {a b} {A : Set a} {B : Set b} →
-    Extensionality (a ⊔ b) (# 0) →
-    Is-proposition A → Is-proposition B →
-    Is-proposition (A Xor B)
-  Xor-closure-propositional {ℓa} {ℓb} {A} {B} ext pA pB =
-    _⇔_.from propositional⇔irrelevant irr
-    where
-    irr : (x y : A Xor B) → x ≡ y
-    irr (inj₁ (a , ¬b)) (inj₂ (¬a  , b))   = ⊥-elim (¬a a)
-    irr (inj₂ (¬a , b)) (inj₁ (a   , ¬b))  = ⊥-elim (¬b b)
-    irr (inj₁ (a , ¬b)) (inj₁ (a′  , ¬b′)) =
-      cong₂ (λ x y → inj₁ (x , y))
-        (_⇔_.to propositional⇔irrelevant pA a a′)
-        (lower-extensionality ℓa _ ext λ b → ⊥-elim (¬b b))
-    irr (inj₂ (¬a , b)) (inj₂ (¬a′ , b′)) =
-      cong₂ (λ x y → inj₂ (x , y))
-        (lower-extensionality ℓb _ ext λ a → ⊥-elim (¬a a))
-        (_⇔_.to propositional⇔irrelevant pB b b′)
-
-  -- However, H-level is not closed under _Xor_.
-
-  ¬-Xor-closure-contractible : ∀ {a b} →
-    ¬ ({A : Set a} {B : Set b} →
-       Contractible A → Contractible B → Contractible (A Xor B))
-  ¬-Xor-closure-contractible closure
-    with proj₁ $ closure (↑-closure 0 ⊤-contractible)
-                         (↑-closure 0 ⊤-contractible)
-  ... | inj₁ (_ , ¬⊤) = ¬⊤ _
-  ... | inj₂ (¬⊤ , _) = ¬⊤ _
-
-  -- Alternative definition of ⊎-closure (for Set₀).
-
-  module Alternative-proof where
-
-    -- Is-set is closed under _⊎_, by an argument similar to the one
-    -- Hedberg used to prove that decidable equality implies
-    -- uniqueness of identity proofs.
-
-    ⊎-closure-set : {A B : Set} →
-                    Is-set A → Is-set B → Is-set (A ⊎ B)
-    ⊎-closure-set {A} {B} A-set B-set =
-      _⇔_.from set⇔UIP (DUIP.constant⇒UIP c)
-      where
-      c : (x y : A ⊎ B) → ∃ λ (f : x ≡ y → x ≡ y) → DUIP.Constant f
-      c (inj₁ x) (inj₁ y) = (cong inj₁ ∘ ⊎.cancel-inj₁ , λ p q → cong (cong inj₁) $ proj₁ $ A-set x y (⊎.cancel-inj₁ p) (⊎.cancel-inj₁ q))
-      c (inj₂ x) (inj₂ y) = (cong inj₂ ∘ ⊎.cancel-inj₂ , λ p q → cong (cong inj₂) $ proj₁ $ B-set x y (⊎.cancel-inj₂ p) (⊎.cancel-inj₂ q))
-      c (inj₁ x) (inj₂ y) = (⊥-elim ∘ ⊎.inj₁≢inj₂       , λ _ → ⊥-elim ∘ ⊎.inj₁≢inj₂)
-      c (inj₂ x) (inj₁ y) = (⊥-elim ∘ ⊎.inj₁≢inj₂ ∘ sym , λ _ → ⊥-elim ∘ ⊎.inj₁≢inj₂ ∘ sym)
-
-    -- H-level is closed under _⊎_ for other levels greater than or equal
-    -- to 2 too.
-
-    ⊎-closure′ :
-      ∀ {A B : Set} n →
-      H-level (2 + n) A → H-level (2 + n) B → H-level (2 + n) (A ⊎ B)
-    ⊎-closure′         zero    = ⊎-closure-set
-    ⊎-closure′ {A} {B} (suc n) = clos
-      where
-      clos : H-level (3 + n) A → H-level (3 + n) B → H-level (3 + n) (A ⊎ B)
-      clos hA hB (inj₁ x) (inj₁ y) = respects-surjection (_↔_.surjection ≡↔inj₁≡inj₁) (2 + n) (hA x y)
-      clos hA hB (inj₂ x) (inj₂ y) = respects-surjection (_↔_.surjection ≡↔inj₂≡inj₂) (2 + n) (hB x y)
-      clos hA hB (inj₁ x) (inj₂ y) = ⊥-elim ∘ ⊎.inj₁≢inj₂
-      clos hA hB (inj₂ x) (inj₁ y) = ⊥-elim ∘ ⊎.inj₁≢inj₂ ∘ sym
+    clos : H-level (3 + n) A → H-level (3 + n) B → H-level (3 + n) (A ⊎ B)
+    clos hA hB (inj₁ x) (inj₁ y) = respects-surjection (_↔_.surjection ≡↔inj₁≡inj₁) (2 + n) (hA x y)
+    clos hA hB (inj₂ x) (inj₂ y) = respects-surjection (_↔_.surjection ≡↔inj₂≡inj₂) (2 + n) (hB x y)
+    clos hA hB (inj₁ x) (inj₂ y) = ⊥-elim ∘ ⊎.inj₁≢inj₂
+    clos hA hB (inj₂ x) (inj₁ y) = ⊥-elim ∘ ⊎.inj₁≢inj₂ ∘ sym
hunk ./H-level/Truncation.agda 94
-abstract
-
-  -- The eliminator gives the right result, up to propositional
-  -- equality, when applied to ∣ x ∣ and ∣ x ∣.
-
-  prop-elim-∣∣ :
-    ∀ {ℓ a p} {A : Set a}
-    (ext : Extensionality (lsuc ℓ ⊔ a) (ℓ ⊔ a ⊔ p))
-    (P : ∥ A ∥ 1 ℓ → Set p)
-    (P-prop : ∀ x → Is-proposition (P x))
-    (f : (x : A) → P ∣ x ∣) (x : A) →
-    prop-elim ext P P-prop f ∣ x ∣ ∣ x ∣ ≡ f x
-  prop-elim-∣∣ _ _ P-prop _ _ =
-    _⇔_.to propositional⇔irrelevant (P-prop _) _ _
+-- The eliminator gives the right result, up to propositional
+-- equality, when applied to ∣ x ∣ and ∣ x ∣.
+
+prop-elim-∣∣ :
+  ∀ {ℓ a p} {A : Set a}
+  (ext : Extensionality (lsuc ℓ ⊔ a) (ℓ ⊔ a ⊔ p))
+  (P : ∥ A ∥ 1 ℓ → Set p)
+  (P-prop : ∀ x → Is-proposition (P x))
+  (f : (x : A) → P ∣ x ∣) (x : A) →
+  prop-elim ext P P-prop f ∣ x ∣ ∣ x ∣ ≡ f x
+prop-elim-∣∣ _ _ P-prop _ _ =
+  _⇔_.to propositional⇔irrelevant (P-prop _) _ _
hunk ./Injection.agda 56
-  abstract
-    injective′ : Injective to′
-    injective′ = injective g ⊚ injective f
+  injective′ : Injective to′
+  injective′ = injective g ⊚ injective f
hunk ./M.agda 54
-abstract
-
-  -- Equality between elements of an M-type can be proved using a pair
-  -- of equalities (assuming extensionality and a kind of η law).
-  --
-  -- Note that M-≡,≡↠≡ does not preserve or establish guardedness, so
-  -- this lemma is perhaps not very useful.
-
-  M-≡,≡↠≡ : ∀ {a b} {A : Set a} {B : A → Set b} →
-            (∀ {x} {C : B x → Set (a ⊔ b)} → Extensionality′ (B x) C) →
-            (∀ {x} {f g : B x → ∞ (M A B)} → ♭ ∘ f ≡ ♭ ∘ g ↠ f ≡ g) →
-            ∀ {x y} {f : B x → ∞ (M A B)} {g : B y → ∞ (M A B)} →
-            (∃ λ (p : x ≡ y) → ∀ i → ♭ (f i) ≡ ♭ (g (subst B p i))) ↠
-            _≡_ {A = M A B} (dns x f) (dns y g)
-  M-≡,≡↠≡ {a} {A = A} {B} ext η {x} {y} {f} {g} =
-    (∃ λ (p : x ≡ y) → ∀ i → ♭ (f i) ≡ ♭ (g (subst B p i)))    ↠⟨ Surjection.∃-cong lemma ⟩
-    (∃ λ (p : x ≡ y) → subst (λ x → B x → ∞ (M A B)) p f ≡ g)  ↠⟨ _↔_.surjection Σ-≡,≡↔≡ ⟩
-    (_≡_ {A = ∃ λ (x : A) → B x → ∞ (M A B)} (x , f) (y , g))  ↠⟨ ↠-≡ (_↔_.surjection (Bijection.inverse (M-unfolding {A = A} {B = B}))) ⟩□
-    (dns x f ≡ dns y g)                                        □
-    where
-    lemma : (p : x ≡ y) →
-            (∀ i → ♭ (f i) ≡ ♭ (g (subst B p i))) ↠
-            (subst (λ x → B x → ∞ (M A B)) p f ≡ g)
-    lemma p = elim
-      (λ {x y} p → (f : B x → ∞ (M A B)) (g : B y → ∞ (M A B)) →
-                   (∀ i → ♭ (f i) ≡ ♭ (g (subst B p i))) ↠
-                   (subst (λ x → B x → ∞ (M A B)) p f ≡ g))
-      (λ x f g →
-         (∀ i → ♭ (f i) ≡ ♭ (g (subst B (refl x) i)))    ↠⟨ subst (λ h → (∀ i → ♭ (f i) ≡ ♭ (g (h i))) ↠ (∀ i → ♭ (f i) ≡ ♭ (g i)))
-                                                                  (sym (lower-extensionality₂ a ext (subst-refl B)))
-                                                                  Surjection.id ⟩
-         (∀ i → ♭ (f i) ≡ ♭ (g i))                       ↠⟨ ext-surj ext ⟩
-         (♭ ∘ f ≡ ♭ ∘ g)                                 ↠⟨ η ⟩
-         (f ≡ g)                                         ↠⟨ subst (λ h → (f ≡ g) ↠ (h ≡ g))
-                                                                  (sym $ subst-refl (λ x' → B x' → ∞ (M A B)) f)
-                                                                  Surjection.id ⟩□
-         (subst (λ x → B x → ∞ (M A B)) (refl x) f ≡ g)  □)
-      p f g
+-- Equality between elements of an M-type can be proved using a pair
+-- of equalities (assuming extensionality and a kind of η law).
+--
+-- Note that M-≡,≡↠≡ does not preserve or establish guardedness, so
+-- this lemma is perhaps not very useful.
+
+M-≡,≡↠≡ : ∀ {a b} {A : Set a} {B : A → Set b} →
+          (∀ {x} {C : B x → Set (a ⊔ b)} → Extensionality′ (B x) C) →
+          (∀ {x} {f g : B x → ∞ (M A B)} → ♭ ∘ f ≡ ♭ ∘ g ↠ f ≡ g) →
+          ∀ {x y} {f : B x → ∞ (M A B)} {g : B y → ∞ (M A B)} →
+          (∃ λ (p : x ≡ y) → ∀ i → ♭ (f i) ≡ ♭ (g (subst B p i))) ↠
+          _≡_ {A = M A B} (dns x f) (dns y g)
+M-≡,≡↠≡ {a} {A = A} {B} ext η {x} {y} {f} {g} =
+  (∃ λ (p : x ≡ y) → ∀ i → ♭ (f i) ≡ ♭ (g (subst B p i)))    ↠⟨ Surjection.∃-cong lemma ⟩
+  (∃ λ (p : x ≡ y) → subst (λ x → B x → ∞ (M A B)) p f ≡ g)  ↠⟨ _↔_.surjection Σ-≡,≡↔≡ ⟩
+  (_≡_ {A = ∃ λ (x : A) → B x → ∞ (M A B)} (x , f) (y , g))  ↠⟨ ↠-≡ (_↔_.surjection (Bijection.inverse (M-unfolding {A = A} {B = B}))) ⟩□
+  (dns x f ≡ dns y g)                                        □
+  where
+  lemma : (p : x ≡ y) →
+          (∀ i → ♭ (f i) ≡ ♭ (g (subst B p i))) ↠
+          (subst (λ x → B x → ∞ (M A B)) p f ≡ g)
+  lemma p = elim
+    (λ {x y} p → (f : B x → ∞ (M A B)) (g : B y → ∞ (M A B)) →
+                 (∀ i → ♭ (f i) ≡ ♭ (g (subst B p i))) ↠
+                 (subst (λ x → B x → ∞ (M A B)) p f ≡ g))
+    (λ x f g →
+       (∀ i → ♭ (f i) ≡ ♭ (g (subst B (refl x) i)))    ↠⟨ subst (λ h → (∀ i → ♭ (f i) ≡ ♭ (g (h i))) ↠ (∀ i → ♭ (f i) ≡ ♭ (g i)))
+                                                                (sym (lower-extensionality₂ a ext (subst-refl B)))
+                                                                Surjection.id ⟩
+       (∀ i → ♭ (f i) ≡ ♭ (g i))                       ↠⟨ ext-surj ext ⟩
+       (♭ ∘ f ≡ ♭ ∘ g)                                 ↠⟨ η ⟩
+       (f ≡ g)                                         ↠⟨ subst (λ h → (f ≡ g) ↠ (h ≡ g))
+                                                                (sym $ subst-refl (λ x' → B x' → ∞ (M A B)) f)
+                                                                Surjection.id ⟩□
+       (subst (λ x → B x → ∞ (M A B)) (refl x) f ≡ g)  □)
+    p f g
hunk ./M.agda 142
-abstract
-
-  -- If we assume a notion of extensionality (bisimilarity implies
-  -- equality) then Contractible is closed under M.
-
-  M-closure-contractible :
-    ∀ {a b} {A : Set a} {B : A → Set b} →
-    ({x y : M A B} → x ≡M y → x ≡ y) →
-    Contractible A → Contractible (M A B)
-  M-closure-contractible {A = A} {B} ext (z , irrA) = (x , ext ∘ irr)
-    where
-    x : M A B
-    x = dns z (λ _ → ♯ x)
-
-    irr : ∀ y → x ≡M y
-    irr (dns x′ f) = dns (irrA x′) λ _ → ♯ irr _
-
-  -- The same applies to Is-proposition.
-
-  M-closure-propositional :
-    ∀ {a b} {A : Set a} {B : A → Set b} →
-    ({x y : M A B} → x ≡M y → x ≡ y) →
-    Is-proposition A → Is-proposition (M A B)
-  M-closure-propositional {A = A} {B} ext p =
-    _⇔_.from propositional⇔irrelevant
-             (λ x y → ext $ irrelevant x y)
-    where
-    irrelevant : (x y : M A B) → x ≡M y
-    irrelevant (dns x f) (dns y g) =
-      dns (proj₁ (p x y)) λ _ → ♯ irrelevant _ _
-
-  -- If we assume that we have another notion of extensionality, then
-  -- Contractible is closed under M.
-
-  M-closure-set :
-    ∀ {a b} {A : Set a} {B : A → Set b} →
-    ({x y : M A B} {p q : x ≡ y} → ≡⇒≡M p ≡≡M ≡⇒≡M q → p ≡ q) →
-    Is-set A → Is-set (M A B)
-  M-closure-set {A = A} {B} ext s =
-    _⇔_.from set⇔UIP (λ p q → ext $ uip (≡⇒≡M p) (≡⇒≡M q))
-    where
-    uip : {x y : M A B} (p q : x ≡M y) → p ≡≡M q
-    uip (dns p _) (dns q _) =
-      dns (proj₁ (s _ _ p q)) λ _ → ♯ uip _ _
+-- If we assume a notion of extensionality (bisimilarity implies
+-- equality) then Contractible is closed under M.
+
+M-closure-contractible :
+  ∀ {a b} {A : Set a} {B : A → Set b} →
+  ({x y : M A B} → x ≡M y → x ≡ y) →
+  Contractible A → Contractible (M A B)
+M-closure-contractible {A = A} {B} ext (z , irrA) = (x , ext ∘ irr)
+  where
+  x : M A B
+  x = dns z (λ _ → ♯ x)
+
+  irr : ∀ y → x ≡M y
+  irr (dns x′ f) = dns (irrA x′) λ _ → ♯ irr _
+
+-- The same applies to Is-proposition.
+
+M-closure-propositional :
+  ∀ {a b} {A : Set a} {B : A → Set b} →
+  ({x y : M A B} → x ≡M y → x ≡ y) →
+  Is-proposition A → Is-proposition (M A B)
+M-closure-propositional {A = A} {B} ext p =
+  _⇔_.from propositional⇔irrelevant
+           (λ x y → ext $ irrelevant x y)
+  where
+  irrelevant : (x y : M A B) → x ≡M y
+  irrelevant (dns x f) (dns y g) =
+    dns (proj₁ (p x y)) λ _ → ♯ irrelevant _ _
+
+-- If we assume that we have another notion of extensionality, then
+-- Contractible is closed under M.
+
+M-closure-set :
+  ∀ {a b} {A : Set a} {B : A → Set b} →
+  ({x y : M A B} {p q : x ≡ y} → ≡⇒≡M p ≡≡M ≡⇒≡M q → p ≡ q) →
+  Is-set A → Is-set (M A B)
+M-closure-set {A = A} {B} ext s =
+  _⇔_.from set⇔UIP (λ p q → ext $ uip (≡⇒≡M p) (≡⇒≡M q))
+  where
+  uip : {x y : M A B} (p q : x ≡M y) → p ≡≡M q
+  uip (dns p _) (dns q _) =
+    dns (proj₁ (s _ _ p q)) λ _ → ♯ uip _ _
hunk ./Preimage.agda 65
-  abstract
-    right-inverse-of : ∀ p → to′ (from′ p) ≡ p
-    right-inverse-of = λ g⁻¹y → cong (_,_ (proj₁ g⁻¹y)) (
-      let p = f≡g (proj₁ g⁻¹y); q = proj₂ g⁻¹y in
-      trans (sym p) (trans p q)  ≡⟨ sym $ trans-assoc _ _ _ ⟩
-      trans (trans (sym p) p) q  ≡⟨ cong (λ p → trans p q) (trans-symˡ _) ⟩
-      trans (refl _) q           ≡⟨ trans-reflˡ _ ⟩∎
-      q                          ∎)
-
-    left-inverse-of : ∀ p → from′ (to′ p) ≡ p
-    left-inverse-of = λ f⁻¹y → cong (_,_ (proj₁ f⁻¹y))
-      let p = f≡g (proj₁ f⁻¹y); q = proj₂ f⁻¹y in
-      trans p (trans (sym p) q)  ≡⟨ sym $ trans-assoc _ _ _ ⟩
-      trans (trans p (sym p)) q  ≡⟨ cong (λ p → trans p q) (trans-symʳ _) ⟩
-      trans (refl _) q           ≡⟨ trans-reflˡ _ ⟩∎
-      q                          ∎
+  right-inverse-of : ∀ p → to′ (from′ p) ≡ p
+  right-inverse-of = λ g⁻¹y → cong (_,_ (proj₁ g⁻¹y)) (
+    let p = f≡g (proj₁ g⁻¹y); q = proj₂ g⁻¹y in
+    trans (sym p) (trans p q)  ≡⟨ sym $ trans-assoc _ _ _ ⟩
+    trans (trans (sym p) p) q  ≡⟨ cong (λ p → trans p q) (trans-symˡ _) ⟩
+    trans (refl _) q           ≡⟨ trans-reflˡ _ ⟩∎
+    q                          ∎)
+
+  left-inverse-of : ∀ p → from′ (to′ p) ≡ p
+  left-inverse-of = λ f⁻¹y → cong (_,_ (proj₁ f⁻¹y))
+    let p = f≡g (proj₁ f⁻¹y); q = proj₂ f⁻¹y in
+    trans p (trans (sym p) q)  ≡⟨ sym $ trans-assoc _ _ _ ⟩
+    trans (trans p (sym p)) q  ≡⟨ cong (λ p → trans p q) (trans-symʳ _) ⟩
+    trans (refl _) q           ≡⟨ trans-reflˡ _ ⟩∎
+    q                          ∎
hunk ./Preimage.agda 105
-  abstract
-    add-∘-lemma : ∀ {x} → from x ≡ y → from (to (from x)) ≡ y
-    add-∘-lemma {x} from-x≡y =
-      from (to (from x))  ≡⟨ cong from (right-inverse-of x) ⟩
-      from x              ≡⟨ from-x≡y ⟩∎
-      y                   ∎
+  add-∘-lemma : ∀ {x} → from x ≡ y → from (to (from x)) ≡ y
+  add-∘-lemma {x} from-x≡y =
+    from (to (from x))  ≡⟨ cong from (right-inverse-of x) ⟩
+    from x              ≡⟨ from-x≡y ⟩∎
+    y                   ∎
hunk ./Preimage.agda 114
-  abstract
-
-    -- add-∘ is a right inverse of drop-∘.
-
-    right-inv : (from⁻¹y : from ⁻¹ y) → drop-∘ (add-∘ from⁻¹y) ≡ from⁻¹y
-    right-inv (x , from-x≡y) =
-        (to (from x) , trans (cong from (right-inverse-of x)) from-x≡y)  ≡⟨ sym $ lemma (right-inverse-of x) from-x≡y ⟩∎
-        (x           , from-x≡y)                                         ∎
-      where
-      lemma : ∀ {x y z} {f : B → A}
-              (y≡x : y ≡ x) (p : f x ≡ z) →
-              _≡_ {A = f ⁻¹ z} (x , p) (y , trans (cong f y≡x) p)
-      lemma {z = z} {f} = elim
-        (λ {y x} y≡x → (p : f x ≡ z) →
-           _≡_ {A = f ⁻¹ z} (x , p) (y , trans (cong f y≡x) p))
-        (λ x p → cong (_,_ x) (
-           p                          ≡⟨ sym $ trans-reflˡ _ ⟩
-           trans (refl (f x)) p       ≡⟨ cong (λ q → trans q p) (sym (cong-refl f)) ⟩∎
-           trans (cong f (refl x)) p  ∎))
+  -- add-∘ is a right inverse of drop-∘.
+
+  right-inv : (from⁻¹y : from ⁻¹ y) → drop-∘ (add-∘ from⁻¹y) ≡ from⁻¹y
+  right-inv (x , from-x≡y) =
+      (to (from x) , trans (cong from (right-inverse-of x)) from-x≡y)  ≡⟨ sym $ lemma (right-inverse-of x) from-x≡y ⟩∎
+      (x           , from-x≡y)                                         ∎
+    where
+    lemma : ∀ {x y z} {f : B → A}
+            (y≡x : y ≡ x) (p : f x ≡ z) →
+            _≡_ {A = f ⁻¹ z} (x , p) (y , trans (cong f y≡x) p)
+    lemma {z = z} {f} = elim
+      (λ {y x} y≡x → (p : f x ≡ z) →
+         _≡_ {A = f ⁻¹ z} (x , p) (y , trans (cong f y≡x) p))
+      (λ x p → cong (_,_ x) (
+         p                          ≡⟨ sym $ trans-reflˡ _ ⟩
+         trans (refl (f x)) p       ≡⟨ cong (λ q → trans q p) (sym (cong-refl f)) ⟩∎
+         trans (cong f (refl x)) p  ∎))
hunk ./Preimage.agda 150
-abstract
-
-  -- Preimages under an injection into a set are propositional.
-
-  injection⁻¹-propositional :
-    ∀ {a b} {A : Set a} {B : Set b} (A↣B : A ↣ B) → let open _↣_ A↣B in
-    Is-set B →
-    ∀ y → Is-proposition (to ⁻¹ y)
-  injection⁻¹-propositional A↣B B-set y =
-    _⇔_.from propositional⇔irrelevant λ { (x₁ , tox₁≡y) (x₂ , tox₂≡y) →
-      Σ-≡,≡→≡ (injective (to x₁  ≡⟨ tox₁≡y ⟩
-                          y      ≡⟨ sym tox₂≡y ⟩∎
-                          to x₂  ∎))
-              (subst (λ x → to x ≡ y)
-                     (injective (trans tox₁≡y (sym tox₂≡y)))
-                     tox₁≡y                                   ≡⟨ _⇔_.to set⇔UIP B-set _ _ ⟩∎
-               tox₂≡y                                         ∎) }
-    where
-    open _↣_ A↣B
+-- Preimages under an injection into a set are propositional.
+
+injection⁻¹-propositional :
+  ∀ {a b} {A : Set a} {B : Set b} (A↣B : A ↣ B) → let open _↣_ A↣B in
+  Is-set B →
+  ∀ y → Is-proposition (to ⁻¹ y)
+injection⁻¹-propositional A↣B B-set y =
+  _⇔_.from propositional⇔irrelevant λ { (x₁ , tox₁≡y) (x₂ , tox₂≡y) →
+    Σ-≡,≡→≡ (injective (to x₁  ≡⟨ tox₁≡y ⟩
+                        y      ≡⟨ sym tox₂≡y ⟩∎
+                        to x₂  ∎))
+            (subst (λ x → to x ≡ y)
+                   (injective (trans tox₁≡y (sym tox₂≡y)))
+                   tox₁≡y                                   ≡⟨ _⇔_.to set⇔UIP B-set _ _ ⟩∎
+             tox₂≡y                                         ∎) }
+  where
+  open _↣_ A↣B
hunk ./Prelude.agda 135
-abstract
-
-  -- Some lemmas.
-
-  zero≤ : ∀ n → zero ≤ n
-  zero≤ zero    = ≤-refl
-  zero≤ (suc n) = ≤-step (zero≤ n)
-
-  suc≤suc : ∀ {m n} → m ≤ n → suc m ≤ suc n
-  suc≤suc ≤-refl       = ≤-refl
-  suc≤suc (≤-step m≤n) = ≤-step (suc≤suc m≤n)
-
-  m≤m+n : ∀ m n → m ≤ m + n
-  m≤m+n zero    n = zero≤ n
-  m≤m+n (suc m) n = suc≤suc (m≤m+n m n)
+-- Some lemmas.
+
+zero≤ : ∀ n → zero ≤ n
+zero≤ zero    = ≤-refl
+zero≤ (suc n) = ≤-step (zero≤ n)
+
+suc≤suc : ∀ {m n} → m ≤ n → suc m ≤ suc n
+suc≤suc ≤-refl       = ≤-refl
+suc≤suc (≤-step m≤n) = ≤-step (suc≤suc m≤n)
+
+m≤m+n : ∀ m n → m ≤ m + n
+m≤m+n zero    n = zero≤ n
+m≤m+n (suc m) n = suc≤suc (m≤m+n m n)
hunk ./Prelude.agda 266
-abstract
-
-  inhabited⇒W-empty : ∀ {a b} {A : Set a} {B : A → Set b} →
-                      (∀ x → B x) → ¬ W A B
-  inhabited⇒W-empty b (sup x f) = inhabited⇒W-empty b (f (b x))
+inhabited⇒W-empty : ∀ {a b} {A : Set a} {B : A → Set b} →
+                    (∀ x → B x) → ¬ W A B
+inhabited⇒W-empty b (sup x f) = inhabited⇒W-empty b (f (b x))
hunk ./Structure-identity-principle.agda 93
-abstract
-
-  structure-identity-principle :
-    ∀ {c₁ c₂ ℓ₁ ℓ₂} →
-    Extensionality (ℓ₁ ⊔ ℓ₂) (ℓ₁ ⊔ ℓ₂) →
-    (C : Category c₁ c₂) →
-    (S : Standard-notion-of-structure ℓ₁ ℓ₂ (Category.precategory C)) →
-    ∀ {X Y} → Is-equivalence
-                (Precategory.≡→≅ (Standard-notion-of-structure.Str S)
-                                 {X} {Y})
-  structure-identity-principle ext C S =
-    Str.≡→≅-equivalence-lemma ≡≃≅ ≡≃≅-refl
-    where
-    open Standard-notion-of-structure S
-    module C = Category C
-
-    -- _≡_ is pointwise equivalent to Str._≅_.
-
-    module ≅HH≃≅ where
-
-      to : ∀ {X Y} {p : P X} {q : P Y} →
-           (∃ λ (f : X C.≅ Y) → H p q (f C.¹) × H q p (f C.⁻¹)) →
-           (X , p) Str.≅ (Y , q)
-      to ((f , f⁻¹ , f∙f⁻¹ , f⁻¹∙f) , Hf , Hf⁻¹) =
-        (f , Hf) , (f⁻¹ , Hf⁻¹) ,
-        lift-equality f∙f⁻¹ ,
-        lift-equality f⁻¹∙f
-
-    ≅HH≃≅ : ∀ {X Y} {p : P X} {q : P Y} →
-            (∃ λ (f : X C.≅ Y) → H p q (f C.¹) × H q p (f C.⁻¹)) ≃
-            ((X , p) Str.≅ (Y , q))
-    ≅HH≃≅ {X} {Y} {p} {q} = ↔⇒≃ (record
-      { surjection = record
-        { logical-equivalence = record
-          { to   = ≅HH≃≅.to
-          ; from = from
-          }
-        ; right-inverse-of = to∘from
+structure-identity-principle :
+  ∀ {c₁ c₂ ℓ₁ ℓ₂} →
+  Extensionality (ℓ₁ ⊔ ℓ₂) (ℓ₁ ⊔ ℓ₂) →
+  (C : Category c₁ c₂) →
+  (S : Standard-notion-of-structure ℓ₁ ℓ₂ (Category.precategory C)) →
+  ∀ {X Y} → Is-equivalence
+              (Precategory.≡→≅ (Standard-notion-of-structure.Str S)
+                               {X} {Y})
+structure-identity-principle ext C S =
+  Str.≡→≅-equivalence-lemma ≡≃≅ ≡≃≅-refl
+  where
+  open Standard-notion-of-structure S
+  module C = Category C
+
+  -- _≡_ is pointwise equivalent to Str._≅_.
+
+  module ≅HH≃≅ where
+
+    to : ∀ {X Y} {p : P X} {q : P Y} →
+         (∃ λ (f : X C.≅ Y) → H p q (f C.¹) × H q p (f C.⁻¹)) →
+         (X , p) Str.≅ (Y , q)
+    to ((f , f⁻¹ , f∙f⁻¹ , f⁻¹∙f) , Hf , Hf⁻¹) =
+      (f , Hf) , (f⁻¹ , Hf⁻¹) ,
+      lift-equality f∙f⁻¹ ,
+      lift-equality f⁻¹∙f
+
+  ≅HH≃≅ : ∀ {X Y} {p : P X} {q : P Y} →
+          (∃ λ (f : X C.≅ Y) → H p q (f C.¹) × H q p (f C.⁻¹)) ≃
+          ((X , p) Str.≅ (Y , q))
+  ≅HH≃≅ {X} {Y} {p} {q} = ↔⇒≃ (record
+    { surjection = record
+      { logical-equivalence = record
+        { to   = ≅HH≃≅.to
+        ; from = from
hunk ./Structure-identity-principle.agda 128
-      ; left-inverse-of = from∘to
-      })
-      where
-      from : (X , p) Str.≅ (Y , q) →
-             ∃ λ (f : X C.≅ Y) → H p q (f C.¹) × H q p (f C.⁻¹)
-      from ((f , Hf) , (f⁻¹ , Hf⁻¹) , f∙f⁻¹ , f⁻¹∙f) =
-        (f , f⁻¹ , cong proj₁ f∙f⁻¹ , cong proj₁ f⁻¹∙f) , Hf , Hf⁻¹
-
-      to∘from : ∀ p → ≅HH≃≅.to (from p) ≡ p
-      to∘from ((f , Hf) , (f⁻¹ , Hf⁻¹) , f∙f⁻¹ , f⁻¹∙f) =
-        cong₂ (λ f∙f⁻¹ f⁻¹∙f →
-                 (f , Hf) , (f⁻¹ , Hf⁻¹) , f∙f⁻¹ , f⁻¹∙f)
-              (_⇔_.to set⇔UIP Str.Hom-is-set _ _)
-              (_⇔_.to set⇔UIP Str.Hom-is-set _ _)
-
-      from∘to : ∀ p → from (≅HH≃≅.to p) ≡ p
-      from∘to ((f , f⁻¹ , f∙f⁻¹ , f⁻¹∙f) , Hf , Hf⁻¹) =
-        cong₂ (λ f∙f⁻¹ f⁻¹∙f → (f , f⁻¹ , f∙f⁻¹ , f⁻¹∙f) , Hf , Hf⁻¹)
-              (_⇔_.to set⇔UIP C.Hom-is-set _ _)
-              (_⇔_.to set⇔UIP C.Hom-is-set _ _)
-
-    module ≡≡≃≅HH where
-
-      to : ∀ {X Y} {p : P X} {q : P Y} →
-           (X≡Y : X ≡ Y) → subst P X≡Y p ≡ q →
-           H p q (C.≡→≅ X≡Y C.¹) × H q p (C.≡→≅ X≡Y C.⁻¹)
-      to {X} {p = p} X≡Y p≡q = elim¹
-        (λ X≡Y → ∀ {q} → subst P X≡Y p ≡ q →
-                 H p q (C.≡→≅ X≡Y C.¹) × H q p (C.≡→≅ X≡Y C.⁻¹))
-        (elim¹
-           (λ {q} _ → H p q (C.≡→≅ (refl X) C.¹) ×
-                      H q p (C.≡→≅ (refl X) C.⁻¹))
-           ( subst (λ { (q , f) → H p q f })
-                   (sym $ cong₂ _,_
-                            (subst P (refl X) p  ≡⟨ subst-refl P _ ⟩∎
-                             p                   ∎)
-                            (C.≡→≅ (refl X) C.¹  ≡⟨ cong C._¹ C.≡→≅-refl ⟩∎
-                             C.id                ∎))
-                   H-id
-           , subst (λ { (q , f) → H q p f })
-                   (sym $ cong₂ _,_
-                            (subst P (refl X) p  ≡⟨ subst-refl P _ ⟩∎
-                             p                   ∎)
-                            (C.≡→≅ (refl X) C.⁻¹  ≡⟨ cong C._⁻¹ C.≡→≅-refl ⟩∎
-                             C.id                 ∎))
-                   H-id
-           ))
-        X≡Y p≡q
-
-      to-refl : ∀ {X} {p : P X} →
-                subst (λ f → H p p (f C.¹) × H p p (f C.⁻¹))
-                      C.≡→≅-refl
-                      (to (refl X) (subst-refl P p)) ≡
-                (H-id , H-id)
-      to-refl =
-        cong₂ _,_ (_⇔_.to propositional⇔irrelevant (H-prop _) _ _)
-                  (_⇔_.to propositional⇔irrelevant (H-prop _) _ _)
-
-    ≡≡≃≅HH : ∀ {X Y} {p : P X} {q : P Y} →
-             (∃ λ (eq : X ≡ Y) → subst P eq p ≡ q) ≃
-             (∃ λ (f : X C.≅ Y) → H p q (f C.¹) × H q p (f C.⁻¹))
-    ≡≡≃≅HH {X} {p = p} {q} =
-      Σ-preserves C.≡≃≅ λ X≡Y →
-        _↔_.to (⇔↔≃ ext (P-set _ _ _)
-                        (×-closure 1 (H-prop _) (H-prop _)))
-               (record { to = ≡≡≃≅HH.to X≡Y ; from = from X≡Y })
-      where
-      from : ∀ X≡Y → H p q (C.≡→≅ X≡Y C.¹) × H q p (C.≡→≅ X≡Y C.⁻¹) →
-             subst P X≡Y p ≡ q
-      from X≡Y (H¹ , H⁻¹) = elim¹
-        (λ {Y} X≡Y → ∀ {q} →
-                     H p q (C.≡→≅ X≡Y C.¹) → H q p (C.≡→≅ X≡Y C.⁻¹) →
-                     subst P X≡Y p ≡ q)
-        (λ {q} H¹ H⁻¹ →
-           subst P (refl X) p  ≡⟨ subst-refl P _ ⟩
-           p                   ≡⟨ H-antisymmetric p q
-                                    (subst (H p q) (cong C._¹  C.≡→≅-refl) H¹)
-                                    (subst (H q p) (cong C._⁻¹ C.≡→≅-refl) H⁻¹) ⟩∎
-           q                   ∎)
-        X≡Y H¹ H⁻¹
-
-    ≡≃≅ : ∀ {X Y} {p : P X} {q : P Y} →
-          _≡_ {A = ∃ P} (X , p) (Y , q) ≃ ((X , p) Str.≅ (Y , q))
-    ≡≃≅ = ≅HH≃≅ ⊚ ≡≡≃≅HH ⊚ ↔⇒≃ (inverse Σ-≡,≡↔≡)
-
-    -- …and the proof maps reflexivity to the identity morphism.
-
-    ≡≃≅-refl : ∀ {Xp} → _≃_.to ≡≃≅ (refl Xp) ≡ Str.id≅
-    ≡≃≅-refl {X , p} =
-      ≅HH≃≅.to (_≃_.to ≡≡≃≅HH (Σ-≡,≡←≡ (refl (_,_ {B = P} X p))))      ≡⟨ cong (≅HH≃≅.to ∘ _≃_.to ≡≡≃≅HH) $ Σ-≡,≡←≡-refl {B = P} ⟩
-      ≅HH≃≅.to (_≃_.to ≡≡≃≅HH (refl X , subst-refl P p))               ≡⟨⟩
-      ≅HH≃≅.to (C.≡→≅ (refl X) , ≡≡≃≅HH.to (refl X) (subst-refl P p))  ≡⟨ cong ≅HH≃≅.to $ Σ-≡,≡→≡ C.≡→≅-refl ≡≡≃≅HH.to-refl ⟩
-      ≅HH≃≅.to (C.id≅ , H-id , H-id)                                   ≡⟨ refl _ ⟩∎
-      Str.id≅                                                          ∎
+      ; right-inverse-of = to∘from
+      }
+    ; left-inverse-of = from∘to
+    })
+    where
+    from : (X , p) Str.≅ (Y , q) →
+           ∃ λ (f : X C.≅ Y) → H p q (f C.¹) × H q p (f C.⁻¹)
+    from ((f , Hf) , (f⁻¹ , Hf⁻¹) , f∙f⁻¹ , f⁻¹∙f) =
+      (f , f⁻¹ , cong proj₁ f∙f⁻¹ , cong proj₁ f⁻¹∙f) , Hf , Hf⁻¹
+
+    to∘from : ∀ p → ≅HH≃≅.to (from p) ≡ p
+    to∘from ((f , Hf) , (f⁻¹ , Hf⁻¹) , f∙f⁻¹ , f⁻¹∙f) =
+      cong₂ (λ f∙f⁻¹ f⁻¹∙f →
+               (f , Hf) , (f⁻¹ , Hf⁻¹) , f∙f⁻¹ , f⁻¹∙f)
+            (_⇔_.to set⇔UIP Str.Hom-is-set _ _)
+            (_⇔_.to set⇔UIP Str.Hom-is-set _ _)
+
+    from∘to : ∀ p → from (≅HH≃≅.to p) ≡ p
+    from∘to ((f , f⁻¹ , f∙f⁻¹ , f⁻¹∙f) , Hf , Hf⁻¹) =
+      cong₂ (λ f∙f⁻¹ f⁻¹∙f → (f , f⁻¹ , f∙f⁻¹ , f⁻¹∙f) , Hf , Hf⁻¹)
+            (_⇔_.to set⇔UIP C.Hom-is-set _ _)
+            (_⇔_.to set⇔UIP C.Hom-is-set _ _)
+
+  module ≡≡≃≅HH where
+
+    to : ∀ {X Y} {p : P X} {q : P Y} →
+         (X≡Y : X ≡ Y) → subst P X≡Y p ≡ q →
+         H p q (C.≡→≅ X≡Y C.¹) × H q p (C.≡→≅ X≡Y C.⁻¹)
+    to {X} {p = p} X≡Y p≡q = elim¹
+      (λ X≡Y → ∀ {q} → subst P X≡Y p ≡ q →
+               H p q (C.≡→≅ X≡Y C.¹) × H q p (C.≡→≅ X≡Y C.⁻¹))
+      (elim¹
+         (λ {q} _ → H p q (C.≡→≅ (refl X) C.¹) ×
+                    H q p (C.≡→≅ (refl X) C.⁻¹))
+         ( subst (λ { (q , f) → H p q f })
+                 (sym $ cong₂ _,_
+                          (subst P (refl X) p  ≡⟨ subst-refl P _ ⟩∎
+                           p                   ∎)
+                          (C.≡→≅ (refl X) C.¹  ≡⟨ cong C._¹ C.≡→≅-refl ⟩∎
+                           C.id                ∎))
+                 H-id
+         , subst (λ { (q , f) → H q p f })
+                 (sym $ cong₂ _,_
+                          (subst P (refl X) p  ≡⟨ subst-refl P _ ⟩∎
+                           p                   ∎)
+                          (C.≡→≅ (refl X) C.⁻¹  ≡⟨ cong C._⁻¹ C.≡→≅-refl ⟩∎
+                           C.id                 ∎))
+                 H-id
+         ))
+      X≡Y p≡q
+
+    to-refl : ∀ {X} {p : P X} →
+              subst (λ f → H p p (f C.¹) × H p p (f C.⁻¹))
+                    C.≡→≅-refl
+                    (to (refl X) (subst-refl P p)) ≡
+              (H-id , H-id)
+    to-refl =
+      cong₂ _,_ (_⇔_.to propositional⇔irrelevant (H-prop _) _ _)
+                (_⇔_.to propositional⇔irrelevant (H-prop _) _ _)
+
+  ≡≡≃≅HH : ∀ {X Y} {p : P X} {q : P Y} →
+           (∃ λ (eq : X ≡ Y) → subst P eq p ≡ q) ≃
+           (∃ λ (f : X C.≅ Y) → H p q (f C.¹) × H q p (f C.⁻¹))
+  ≡≡≃≅HH {X} {p = p} {q} =
+    Σ-preserves C.≡≃≅ λ X≡Y →
+      _↔_.to (⇔↔≃ ext (P-set _ _ _)
+                      (×-closure 1 (H-prop _) (H-prop _)))
+             (record { to = ≡≡≃≅HH.to X≡Y ; from = from X≡Y })
+    where
+    from : ∀ X≡Y → H p q (C.≡→≅ X≡Y C.¹) × H q p (C.≡→≅ X≡Y C.⁻¹) →
+           subst P X≡Y p ≡ q
+    from X≡Y (H¹ , H⁻¹) = elim¹
+      (λ {Y} X≡Y → ∀ {q} →
+                   H p q (C.≡→≅ X≡Y C.¹) → H q p (C.≡→≅ X≡Y C.⁻¹) →
+                   subst P X≡Y p ≡ q)
+      (λ {q} H¹ H⁻¹ →
+         subst P (refl X) p  ≡⟨ subst-refl P _ ⟩
+         p                   ≡⟨ H-antisymmetric p q
+                                  (subst (H p q) (cong C._¹  C.≡→≅-refl) H¹)
+                                  (subst (H q p) (cong C._⁻¹ C.≡→≅-refl) H⁻¹) ⟩∎
+         q                   ∎)
+      X≡Y H¹ H⁻¹
+
+  ≡≃≅ : ∀ {X Y} {p : P X} {q : P Y} →
+        _≡_ {A = ∃ P} (X , p) (Y , q) ≃ ((X , p) Str.≅ (Y , q))
+  ≡≃≅ = ≅HH≃≅ ⊚ ≡≡≃≅HH ⊚ ↔⇒≃ (inverse Σ-≡,≡↔≡)
+
+  -- …and the proof maps reflexivity to the identity morphism.
+
+  ≡≃≅-refl : ∀ {Xp} → _≃_.to ≡≃≅ (refl Xp) ≡ Str.id≅
+  ≡≃≅-refl {X , p} =
+    ≅HH≃≅.to (_≃_.to ≡≡≃≅HH (Σ-≡,≡←≡ (refl (_,_ {B = P} X p))))      ≡⟨ cong (≅HH≃≅.to ∘ _≃_.to ≡≡≃≅HH) $ Σ-≡,≡←≡-refl {B = P} ⟩
+    ≅HH≃≅.to (_≃_.to ≡≡≃≅HH (refl X , subst-refl P p))               ≡⟨⟩
+    ≅HH≃≅.to (C.≡→≅ (refl X) , ≡≡≃≅HH.to (refl X) (subst-refl P p))  ≡⟨ cong ≅HH≃≅.to $ Σ-≡,≡→≡ C.≡→≅-refl ≡≡≃≅HH.to-refl ⟩
+    ≅HH≃≅.to (C.id≅ , H-id , H-id)                                   ≡⟨ refl _ ⟩∎
+    Str.id≅                                                          ∎
hunk ./Structure-identity-principle.agda 299
-  abstract
-
-    -- Every Str morphism is an isomorphism.
-
-    Str-homs-are-isos :
-      ∀ {C D x y} (f : ∃ (H {X = C} {Y = D} x y)) →
-      Str.Is-isomorphism {X = C , x} {Y = D , y} f
-    Str-homs-are-isos {C} {D} {x} {y} (C≅D , x≅y) =
-
-      (D≅C ,
-       (resp a (≅⇒≃ D C D≅C) y            ≡⟨ cong (λ eq → resp a eq y) $ Eq.lift-equality ext (refl _) ⟩
-        resp a (inverse $ ≅⇒≃ C D C≅D) y  ≡⟨ resp-preserves-inverses (El a) (resp a) (resp-id ass a) univ₁ ext (≅⇒≃ C D C≅D) _ _ x≅y ⟩∎
-        x                                 ∎)) ,
-
-      S.lift-equality {X = C , x} {Y = C , x} (
-        C≅D Fun.∙≅ D≅C   ≡⟨ _≃_.from (Fun.≡≃≡¹ {X = C} {Y = C}) (Fun._¹⁻¹ {X = C} {Y = D} C≅D) ⟩∎
-        Fun.id≅ {X = C}  ∎) ,
-
-      S.lift-equality {X = D , y} {Y = D , y} (
-        D≅C Fun.∙≅ C≅D   ≡⟨ _≃_.from (Fun.≡≃≡¹ {X = D} {Y = D}) (Fun._⁻¹¹ {X = C} {Y = D} C≅D) ⟩∎
-        Fun.id≅ {X = D}  ∎)
-
-      where
-
-      D≅C : D Fun.≅ C
-      D≅C = Fun._⁻¹≅ {X = C} {Y = D} C≅D
-
-    -- The isomorphism that should be constructed.
-
-    isomorphic : Isomorphic (a , P) (C , x , p) (D , y , q) ↔
-                 ((C , x , p) ≡ (D , y , q))
-    isomorphic =
-      Σ (C ≃ D) (λ eq → Is-isomorphism a eq x y)     ↝⟨ (let ≃≃≅-CD = ≃≃≅-Set (# 1) ext (C , C-set) (D , D-set) in
-                                                         Σ-cong ≃≃≅-CD (λ eq →
-                                                           let eq′ = _≃_.from ≃≃≅-CD (_≃_.to ≃≃≅-CD eq) in
-                                                           Is-isomorphism a eq  x y  ↝⟨ ≡⇒↝ _ $ cong (λ eq → Is-isomorphism a eq x y) $ sym $
-                                                                                          _≃_.left-inverse-of ≃≃≅-CD eq ⟩
-                                                           Is-isomorphism a eq′ x y  □)) ⟩
-      ∃ (H {X = C , C-set} {Y = D , D-set} x y)      ↝⟨ inverse ×-right-identity ⟩
-      ∃ (H {X = C , C-set} {Y = D , D-set} x y) × ⊤  ↝⟨ ∃-cong (λ X≅Y → _⇔_.to contractible⇔⊤↔ $ propositional⇒inhabited⇒contractible
-                                                                          (Str.Is-isomorphism-propositional X≅Y)
-                                                                          (Str-homs-are-isos X≅Y)) ⟩
-      (((C , C-set) , x) Str.≅ ((D , D-set) , y))    ↔⟨ inverse ⟨ _ , structure-identity-principle ext Bij S
-                                                                        {X = (C , C-set) , x} {Y = (D , D-set) , y} ⟩ ⟩
-      (((C , C-set) , x) ≡ ((D , D-set) , y))        ↔⟨ ≃-≡ $ ↔⇒≃ (Σ-assoc ⊚ ∃-cong (λ _ → ×-comm) ⊚ inverse Σ-assoc) ⟩
-      (((C , x) , C-set) ≡ ((D , y) , D-set))        ↝⟨ inverse $ ignore-propositional-component (H-level-propositional ext 2) ⟩
-      ((C , x) ≡ (D , y))                            ↝⟨ ignore-propositional-component (proj₂ (P D y) ass) ⟩
-      (((C , x) , p) ≡ ((D , y) , q))                ↔⟨ ≃-≡ $ ↔⇒≃ Σ-assoc ⟩□
-      ((C , x , p) ≡ (D , y , q))                    □
-
-    -- A simplification lemma.
-
-    proj₁-from-isomorphic :
-      ∀ X≡Y →
-      proj₁ (_↔_.from isomorphic X≡Y) ≡
-      elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) X≡Y
-    proj₁-from-isomorphic X≡Y = Eq.lift-equality ext (
-
-      _≃_.to (proj₁ (_↔_.from isomorphic X≡Y))                         ≡⟨⟩
-
-      cast C-set D-set X≡Y                                             ≡⟨ lemma ⟩∎
-
-      _≃_.to (elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) X≡Y)  ∎)
-
-      where
-
-      cast : ∀ {X Y} →
-             Is-set (Carrier (a , P) X) → Is-set (Carrier (a , P) Y) →
-             X ≡ Y → Carrier (a , P) X → Carrier (a , P) Y
-      cast {C , x , p} {D , y , q} C-set D-set X≡Y =
-        proj₁ $ proj₁ $ proj₁ $
-        Str.≡→≅ {X = (C , C-set) , x} {Y = (D , D-set) , y} $
-        cong (λ { ((C , x) , C-set) → (C , C-set) , x }) $
-        Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡
-                   (cong (λ { (C , (x , p)) → (C , x) , p }) X≡Y)))
-                (proj₁ (H-level-propositional ext 2 _ _))
-
-      lemma :
-        cast C-set D-set X≡Y ≡
-        _≃_.to (elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) X≡Y)
-      lemma = elim¹
-        (λ X≡Y → ∀ D-set → cast C-set D-set X≡Y ≡
-                           _≃_.to (elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y)
-                                        (λ _ → Eq.id) X≡Y))
-        (λ C-set′ →
-
-           cast C-set C-set′ (refl (C , x , p))                      ≡⟨ cong (λ eq →
-                                                                                proj₁ $ proj₁ $ proj₁ $
-                                                                                Str.≡→≅ {X = (C , C-set) , x} {Y = (C , C-set′) , x} $
-                                                                                cong (λ { ((C , x) , C-set) → (C , C-set) , x }) $
-                                                                                Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ {B = proj₁ ∘ uncurry P} eq))
-                                                                                        (proj₁ (H-level-propositional ext 2 _ _)))
-                                                                             (cong-refl (λ { (C , (x , p)) → (C , x) , p })) ⟩
-           (proj₁ $ proj₁ $ proj₁ $
-            Str.≡→≅ {X = (C , C-set) , x} {Y = (C , C-set′) , x} $
-            cong (λ { ((C , x) , C-set) → (C , C-set) , x }) $
-            Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
-                    (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong proj₁ (S.proj₁-≡→≅-¹ _) ⟩
-
-           (proj₁ $
-            Fun.≡→≅ $
-            cong proj₁ $
-            cong (λ { ((C , x) , C-set) → (C , C-set) , x }) $
-            Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
-                    (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong (proj₁ ∘ Fun.≡→≅) $
-                                                                          cong-∘ proj₁ (λ { ((C , x) , C-set) → (C , C-set) , x }) _ ⟩
-           (proj₁ $
-            Fun.≡→≅ $
-            cong (λ { ((C , _) , C-set) → C , C-set }) $
-            Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
-                    (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ Fun.≡→≅-¹ _ ⟩
-
-           (elim (λ {X Y} _ → Fun.Hom X Y) (λ _ → P.id) $
-            cong (λ { ((C , _) , C-set) → C , C-set }) $
-            Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
-                    (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ elim¹ (λ eq → elim (λ {X Y} _ → Fun.Hom X Y) (λ _ → P.id) eq ≡
-                                                                                      ≡⇒↝ implication (cong proj₁ eq))
-                        (elim (λ {X Y} _ → Fun.Hom X Y) (λ _ → P.id)
-                              (refl (C , C-set))                          ≡⟨ elim-refl (λ {X Y} _ → Fun.Hom X Y) _ ⟩
-                         P.id                                             ≡⟨ sym $ elim-refl (λ {A B} _ → A → B) _ ⟩
-                         ≡⇒↝ implication (refl C)                         ≡⟨ cong (≡⇒↝ implication) (sym $ cong-refl proj₁) ⟩∎
-                         ≡⇒↝ implication (cong proj₁ (refl (C , C-set)))  ∎) _ ⟩
-
-           (≡⇒↝ implication $
-            cong proj₁ $
-            cong (λ { ((C , _) , C-set) → C , C-set }) $
-            Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
-                    (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong (≡⇒↝ implication)
-                                                                             (cong-∘ proj₁ (λ { ((C , _) , C-set) → C , C-set }) _) ⟩
-           (≡⇒↝ implication $
-            cong (λ { ((C , _) , _) → C }) $
-            Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
-                    (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong (λ eq →
-                                                                                ≡⇒↝ implication $
-                                                                                cong (λ { ((C , _) , _) → C }) $
-                                                                                Σ-≡,≡→≡ (proj₁ {B = λ q → subst (proj₁ ∘ uncurry P) q p ≡ p} eq)
-                                                                                        (proj₁ (H-level-propositional ext 2 _ C-set′)))
-                                                                             (Σ-≡,≡←≡-refl {B = λ { (C , x) → proj₁ (P C x)}}) ⟩
-           (≡⇒↝ implication $
-            cong (λ { ((C , _) , _) → C }) $
-            Σ-≡,≡→≡ (refl (C , x))
-                    (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong (≡⇒↝ implication ∘ cong (λ { ((C , _) , _) → C }))
-                                                                             (Σ-≡,≡→≡-reflˡ (proj₁ (H-level-propositional ext 2 _ _))) ⟩
-           (≡⇒↝ implication $
-            cong (λ { ((C , _) , _) → C }) $
-            cong (_,_ (C , x))
-                 (trans (sym $ subst-refl (Is-set ∘ proj₁) C-set)
-                        (proj₁ (H-level-propositional ext 2 _ _))))  ≡⟨ cong (≡⇒↝ implication)
-                                                                             (cong-∘ (λ { ((C , _) , _) → C }) (_,_ (C , x)) _) ⟩
-           (≡⇒↝ implication $
-            cong (const C)
-                 (trans (sym $ subst-refl (Is-set ∘ proj₁) C-set)
-                        (proj₁ (H-level-propositional ext 2 _ _))))  ≡⟨ cong (≡⇒↝ implication) (cong-const _) ⟩
-
-           ≡⇒↝ implication (refl C)                                  ≡⟨ ≡⇒↝-refl ⟩
-
-           P.id                                                      ≡⟨ sym $ cong _≃_.to $ elim-refl (λ {X Y} _ → proj₁ X ≃ proj₁ Y) _ ⟩∎
-
-           _≃_.to (elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y)
-                        (λ _ → Eq.id)
-                        (refl (C , x , p)))                          ∎)
-
-        X≡Y D-set
-
-abstract
-
-  -- The from component of isomorphism-is-equality′ is equal
-  -- to a simple function.
-
-  from-isomorphism-is-equality′ :
-    ∀ Univ ass c X Y → let open Universe Univ; open Class Univ in
-    (El-set : ∀ {B} → Is-set B → Is-set (El (proj₁ c) B)) →
-    ∀ I-set J-set →
-    _↔_.from (isomorphism-is-equality′
-                Univ ass c X Y El-set I-set J-set) ≡
-    elim (λ {X Y} _ → Isomorphic c X Y)
-         (λ { (_ , x , _) → Eq.id , resp-id ass (proj₁ c) x })
-  from-isomorphism-is-equality′ Univ ass c X Y El-set I-set J-set =
-    ext λ eq →
-      Σ-≡,≡→≡ (lemma eq) (_⇔_.to set⇔UIP (El-set J-set) _ _)
+  -- Every Str morphism is an isomorphism.
+
+  Str-homs-are-isos :
+    ∀ {C D x y} (f : ∃ (H {X = C} {Y = D} x y)) →
+    Str.Is-isomorphism {X = C , x} {Y = D , y} f
+  Str-homs-are-isos {C} {D} {x} {y} (C≅D , x≅y) =
+
+    (D≅C ,
+     (resp a (≅⇒≃ D C D≅C) y            ≡⟨ cong (λ eq → resp a eq y) $ Eq.lift-equality ext (refl _) ⟩
+      resp a (inverse $ ≅⇒≃ C D C≅D) y  ≡⟨ resp-preserves-inverses (El a) (resp a) (resp-id ass a) univ₁ ext (≅⇒≃ C D C≅D) _ _ x≅y ⟩∎
+      x                                 ∎)) ,
+
+    S.lift-equality {X = C , x} {Y = C , x} (
+      C≅D Fun.∙≅ D≅C   ≡⟨ _≃_.from (Fun.≡≃≡¹ {X = C} {Y = C}) (Fun._¹⁻¹ {X = C} {Y = D} C≅D) ⟩∎
+      Fun.id≅ {X = C}  ∎) ,
+
+    S.lift-equality {X = D , y} {Y = D , y} (
+      D≅C Fun.∙≅ C≅D   ≡⟨ _≃_.from (Fun.≡≃≡¹ {X = D} {Y = D}) (Fun._⁻¹¹ {X = C} {Y = D} C≅D) ⟩∎
+      Fun.id≅ {X = D}  ∎)
+
hunk ./Structure-identity-principle.agda 320
-    open Assumptions ass
-    open Universe Univ
-    open Class Univ
+
+    D≅C : D Fun.≅ C
+    D≅C = Fun._⁻¹≅ {X = C} {Y = D} C≅D
+
+  -- The isomorphism that should be constructed.
+
+  isomorphic : Isomorphic (a , P) (C , x , p) (D , y , q) ↔
+               ((C , x , p) ≡ (D , y , q))
+  isomorphic =
+    Σ (C ≃ D) (λ eq → Is-isomorphism a eq x y)     ↝⟨ (let ≃≃≅-CD = ≃≃≅-Set (# 1) ext (C , C-set) (D , D-set) in
+                                                       Σ-cong ≃≃≅-CD (λ eq →
+                                                         let eq′ = _≃_.from ≃≃≅-CD (_≃_.to ≃≃≅-CD eq) in
+                                                         Is-isomorphism a eq  x y  ↝⟨ ≡⇒↝ _ $ cong (λ eq → Is-isomorphism a eq x y) $ sym $
+                                                                                        _≃_.left-inverse-of ≃≃≅-CD eq ⟩
+                                                         Is-isomorphism a eq′ x y  □)) ⟩
+    ∃ (H {X = C , C-set} {Y = D , D-set} x y)      ↝⟨ inverse ×-right-identity ⟩
+    ∃ (H {X = C , C-set} {Y = D , D-set} x y) × ⊤  ↝⟨ ∃-cong (λ X≅Y → _⇔_.to contractible⇔⊤↔ $ propositional⇒inhabited⇒contractible
+                                                                        (Str.Is-isomorphism-propositional X≅Y)
+                                                                        (Str-homs-are-isos X≅Y)) ⟩
+    (((C , C-set) , x) Str.≅ ((D , D-set) , y))    ↔⟨ inverse ⟨ _ , structure-identity-principle ext Bij S
+                                                                      {X = (C , C-set) , x} {Y = (D , D-set) , y} ⟩ ⟩
+    (((C , C-set) , x) ≡ ((D , D-set) , y))        ↔⟨ ≃-≡ $ ↔⇒≃ (Σ-assoc ⊚ ∃-cong (λ _ → ×-comm) ⊚ inverse Σ-assoc) ⟩
+    (((C , x) , C-set) ≡ ((D , y) , D-set))        ↝⟨ inverse $ ignore-propositional-component (H-level-propositional ext 2) ⟩
+    ((C , x) ≡ (D , y))                            ↝⟨ ignore-propositional-component (proj₂ (P D y) ass) ⟩
+    (((C , x) , p) ≡ ((D , y) , q))                ↔⟨ ≃-≡ $ ↔⇒≃ Σ-assoc ⟩□
+    ((C , x , p) ≡ (D , y , q))                    □
+
+  -- A simplification lemma.
+
+  proj₁-from-isomorphic :
+    ∀ X≡Y →
+    proj₁ (_↔_.from isomorphic X≡Y) ≡
+    elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) X≡Y
+  proj₁-from-isomorphic X≡Y = Eq.lift-equality ext (
+
+    _≃_.to (proj₁ (_↔_.from isomorphic X≡Y))                         ≡⟨⟩
+
+    cast C-set D-set X≡Y                                             ≡⟨ lemma ⟩∎
+
+    _≃_.to (elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) X≡Y)  ∎)
+
+    where
+
+    cast : ∀ {X Y} →
+           Is-set (Carrier (a , P) X) → Is-set (Carrier (a , P) Y) →
+           X ≡ Y → Carrier (a , P) X → Carrier (a , P) Y
+    cast {C , x , p} {D , y , q} C-set D-set X≡Y =
+      proj₁ $ proj₁ $ proj₁ $
+      Str.≡→≅ {X = (C , C-set) , x} {Y = (D , D-set) , y} $
+      cong (λ { ((C , x) , C-set) → (C , C-set) , x }) $
+      Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡
+                 (cong (λ { (C , (x , p)) → (C , x) , p }) X≡Y)))
+              (proj₁ (H-level-propositional ext 2 _ _))
hunk ./Structure-identity-principle.agda 375
-      ∀ eq →
-      proj₁ (_↔_.from (isomorphism-is-equality′
-                         Univ ass c X Y El-set I-set J-set) eq) ≡
-      proj₁ (elim (λ {X Y} _ → Isomorphic c X Y)
-                  (λ { (_ , x , _) → Eq.id , resp-id ass (proj₁ c) x })
-                  eq)
-    lemma eq =
-      proj₁ (_↔_.from (isomorphism-is-equality′
-                         Univ ass c X Y El-set I-set J-set) eq)          ≡⟨ Isomorphism-is-equality′.proj₁-from-isomorphic
-                                                                              Univ ass _ _ _ _ _ _ _ _ El-set I-set J-set eq ⟩
-      elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) eq              ≡⟨ sym $ elim-∘ (λ {X Y} _ → Isomorphic c X Y) _ proj₁ _ ⟩∎
-      proj₁ (elim (λ {X Y} _ → Isomorphic c X Y)
-                  (λ { (_ , x , _) → Eq.id , resp-id ass (proj₁ c) x })
-                  eq)                                                    ∎
+      cast C-set D-set X≡Y ≡
+      _≃_.to (elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) X≡Y)
+    lemma = elim¹
+      (λ X≡Y → ∀ D-set → cast C-set D-set X≡Y ≡
+                         _≃_.to (elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y)
+                                      (λ _ → Eq.id) X≡Y))
+      (λ C-set′ →
+
+         cast C-set C-set′ (refl (C , x , p))                      ≡⟨ cong (λ eq →
+                                                                              proj₁ $ proj₁ $ proj₁ $
+                                                                              Str.≡→≅ {X = (C , C-set) , x} {Y = (C , C-set′) , x} $
+                                                                              cong (λ { ((C , x) , C-set) → (C , C-set) , x }) $
+                                                                              Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ {B = proj₁ ∘ uncurry P} eq))
+                                                                                      (proj₁ (H-level-propositional ext 2 _ _)))
+                                                                           (cong-refl (λ { (C , (x , p)) → (C , x) , p })) ⟩
+         (proj₁ $ proj₁ $ proj₁ $
+          Str.≡→≅ {X = (C , C-set) , x} {Y = (C , C-set′) , x} $
+          cong (λ { ((C , x) , C-set) → (C , C-set) , x }) $
+          Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
+                  (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong proj₁ (S.proj₁-≡→≅-¹ _) ⟩
+
+         (proj₁ $
+          Fun.≡→≅ $
+          cong proj₁ $
+          cong (λ { ((C , x) , C-set) → (C , C-set) , x }) $
+          Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
+                  (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong (proj₁ ∘ Fun.≡→≅) $
+                                                                        cong-∘ proj₁ (λ { ((C , x) , C-set) → (C , C-set) , x }) _ ⟩
+         (proj₁ $
+          Fun.≡→≅ $
+          cong (λ { ((C , _) , C-set) → C , C-set }) $
+          Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
+                  (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ Fun.≡→≅-¹ _ ⟩
+
+         (elim (λ {X Y} _ → Fun.Hom X Y) (λ _ → P.id) $
+          cong (λ { ((C , _) , C-set) → C , C-set }) $
+          Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
+                  (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ elim¹ (λ eq → elim (λ {X Y} _ → Fun.Hom X Y) (λ _ → P.id) eq ≡
+                                                                                    ≡⇒↝ implication (cong proj₁ eq))
+                      (elim (λ {X Y} _ → Fun.Hom X Y) (λ _ → P.id)
+                            (refl (C , C-set))                          ≡⟨ elim-refl (λ {X Y} _ → Fun.Hom X Y) _ ⟩
+                       P.id                                             ≡⟨ sym $ elim-refl (λ {A B} _ → A → B) _ ⟩
+                       ≡⇒↝ implication (refl C)                         ≡⟨ cong (≡⇒↝ implication) (sym $ cong-refl proj₁) ⟩∎
+                       ≡⇒↝ implication (cong proj₁ (refl (C , C-set)))  ∎) _ ⟩
+
+         (≡⇒↝ implication $
+          cong proj₁ $
+          cong (λ { ((C , _) , C-set) → C , C-set }) $
+          Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
+                  (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong (≡⇒↝ implication)
+                                                                           (cong-∘ proj₁ (λ { ((C , _) , C-set) → C , C-set }) _) ⟩
+         (≡⇒↝ implication $
+          cong (λ { ((C , _) , _) → C }) $
+          Σ-≡,≡→≡ (proj₁ (Σ-≡,≡←≡ (refl ((C , x) , p))))
+                  (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong (λ eq →
+                                                                              ≡⇒↝ implication $
+                                                                              cong (λ { ((C , _) , _) → C }) $
+                                                                              Σ-≡,≡→≡ (proj₁ {B = λ q → subst (proj₁ ∘ uncurry P) q p ≡ p} eq)
+                                                                                      (proj₁ (H-level-propositional ext 2 _ C-set′)))
+                                                                           (Σ-≡,≡←≡-refl {B = λ { (C , x) → proj₁ (P C x)}}) ⟩
+         (≡⇒↝ implication $
+          cong (λ { ((C , _) , _) → C }) $
+          Σ-≡,≡→≡ (refl (C , x))
+                  (proj₁ (H-level-propositional ext 2 _ _)))       ≡⟨ cong (≡⇒↝ implication ∘ cong (λ { ((C , _) , _) → C }))
+                                                                           (Σ-≡,≡→≡-reflˡ (proj₁ (H-level-propositional ext 2 _ _))) ⟩
+         (≡⇒↝ implication $
+          cong (λ { ((C , _) , _) → C }) $
+          cong (_,_ (C , x))
+               (trans (sym $ subst-refl (Is-set ∘ proj₁) C-set)
+                      (proj₁ (H-level-propositional ext 2 _ _))))  ≡⟨ cong (≡⇒↝ implication)
+                                                                           (cong-∘ (λ { ((C , _) , _) → C }) (_,_ (C , x)) _) ⟩
+         (≡⇒↝ implication $
+          cong (const C)
+               (trans (sym $ subst-refl (Is-set ∘ proj₁) C-set)
+                      (proj₁ (H-level-propositional ext 2 _ _))))  ≡⟨ cong (≡⇒↝ implication) (cong-const _) ⟩
+
+         ≡⇒↝ implication (refl C)                                  ≡⟨ ≡⇒↝-refl ⟩
+
+         P.id                                                      ≡⟨ sym $ cong _≃_.to $ elim-refl (λ {X Y} _ → proj₁ X ≃ proj₁ Y) _ ⟩∎
+
+         _≃_.to (elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y)
+                      (λ _ → Eq.id)
+                      (refl (C , x , p)))                          ∎)
+
+      X≡Y D-set
+
+-- The from component of isomorphism-is-equality′ is equal
+-- to a simple function.
+
+from-isomorphism-is-equality′ :
+  ∀ Univ ass c X Y → let open Universe Univ; open Class Univ in
+  (El-set : ∀ {B} → Is-set B → Is-set (El (proj₁ c) B)) →
+  ∀ I-set J-set →
+  _↔_.from (isomorphism-is-equality′
+              Univ ass c X Y El-set I-set J-set) ≡
+  elim (λ {X Y} _ → Isomorphic c X Y)
+       (λ { (_ , x , _) → Eq.id , resp-id ass (proj₁ c) x })
+from-isomorphism-is-equality′ Univ ass c X Y El-set I-set J-set =
+  ext λ eq →
+    Σ-≡,≡→≡ (lemma eq) (_⇔_.to set⇔UIP (El-set J-set) _ _)
+  where
+  open Assumptions ass
+  open Universe Univ
+  open Class Univ
+
+  lemma :
+    ∀ eq →
+    proj₁ (_↔_.from (isomorphism-is-equality′
+                       Univ ass c X Y El-set I-set J-set) eq) ≡
+    proj₁ (elim (λ {X Y} _ → Isomorphic c X Y)
+                (λ { (_ , x , _) → Eq.id , resp-id ass (proj₁ c) x })
+                eq)
+  lemma eq =
+    proj₁ (_↔_.from (isomorphism-is-equality′
+                       Univ ass c X Y El-set I-set J-set) eq)          ≡⟨ Isomorphism-is-equality′.proj₁-from-isomorphic
+                                                                            Univ ass _ _ _ _ _ _ _ _ El-set I-set J-set eq ⟩
+    elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) eq              ≡⟨ sym $ elim-∘ (λ {X Y} _ → Isomorphic c X Y) _ proj₁ _ ⟩∎
+    proj₁ (elim (λ {X Y} _ → Isomorphic c X Y)
+                (λ { (_ , x , _) → Eq.id , resp-id ass (proj₁ c) x })
+                eq)                                                    ∎
hunk ./Surjection.agda 66
-  abstract
-    to∘from : ∀ x → to f (to g (from g (from f x))) ≡ x
-    to∘from = λ x →
-      to f (to g (from g (from f x)))  ≡⟨ cong (to f) (right-inverse-of g (from f x)) ⟩
-      to f (from f x)                  ≡⟨ right-inverse-of f x ⟩∎
-      x                                ∎
+  to∘from : ∀ x → to f (to g (from g (from f x))) ≡ x
+  to∘from = λ x →
+    to f (to g (from g (from f x)))  ≡⟨ cong (to f) (right-inverse-of g (from f x)) ⟩
+    to f (from f x)                  ≡⟨ right-inverse-of f x ⟩∎
+    x                                ∎
hunk ./Surjection.agda 109
-  abstract
-    right-inverse-of′ : ∀ p → to′ (from′ p) ≡ p
-    right-inverse-of′ = λ p →
-      cong (_,_ (proj₁ p)) (right-inverse-of (B₁↠B₂ (proj₁ p)) _)
+  right-inverse-of′ : ∀ p → to′ (from′ p) ≡ p
+  right-inverse-of′ = λ p →
+    cong (_,_ (proj₁ p)) (right-inverse-of (B₁↠B₂ (proj₁ p)) _)
hunk ./Surjection.agda 134
-  abstract
-    right-inverse-of′ : ∀ p → to′ (cong from p) ≡ p
-    right-inverse-of′ = elim
-      (λ {x y} x≡y → trans (sym (right-inverse-of x)) (
-                       trans (cong to (cong from x≡y)) (
-                       right-inverse-of y)) ≡
-                     x≡y)
-      (λ x → trans (sym (right-inverse-of x)) (
-               trans (cong to (cong from (refl x))) (
-               right-inverse-of x))                                 ≡⟨ cong (λ p → trans (sym (right-inverse-of x))
-                                                                                         (trans (cong to p) (right-inverse-of x)))
-                                                                            (cong-refl from) ⟩
-             trans (sym (right-inverse-of x)) (
-               trans (cong to (refl (from x))) (
-               right-inverse-of x))                                 ≡⟨ cong (λ p → trans (sym (right-inverse-of x))
-                                                                                         (trans p (right-inverse-of x)))
-                                                                            (cong-refl to) ⟩
-             trans (sym (right-inverse-of x)) (
-               trans (refl (to (from x))) (
-               right-inverse-of x))                                 ≡⟨ cong (trans (sym _)) (trans-reflˡ _) ⟩
-             trans (sym (right-inverse-of x)) (right-inverse-of x)  ≡⟨ trans-symˡ _ ⟩∎
-             refl x                                                 ∎)
+  right-inverse-of′ : ∀ p → to′ (cong from p) ≡ p
+  right-inverse-of′ = elim
+    (λ {x y} x≡y → trans (sym (right-inverse-of x)) (
+                     trans (cong to (cong from x≡y)) (
+                     right-inverse-of y)) ≡
+                   x≡y)
+    (λ x → trans (sym (right-inverse-of x)) (
+             trans (cong to (cong from (refl x))) (
+             right-inverse-of x))                                 ≡⟨ cong (λ p → trans (sym (right-inverse-of x))
+                                                                                       (trans (cong to p) (right-inverse-of x)))
+                                                                          (cong-refl from) ⟩
+           trans (sym (right-inverse-of x)) (
+             trans (cong to (refl (from x))) (
+             right-inverse-of x))                                 ≡⟨ cong (λ p → trans (sym (right-inverse-of x))
+                                                                                       (trans p (right-inverse-of x)))
+                                                                          (cong-refl to) ⟩
+           trans (sym (right-inverse-of x)) (
+             trans (refl (to (from x))) (
+             right-inverse-of x))                                 ≡⟨ cong (trans (sym _)) (trans-reflˡ _) ⟩
+           trans (sym (right-inverse-of x)) (right-inverse-of x)  ≡⟨ trans-symˡ _ ⟩∎
+           refl x                                                 ∎)
hunk ./Univalence-axiom.agda 79
-abstract
-
-  -- "Evaluation rule" for ≡⇒≃.
-
-  ≡⇒≃-refl : ∀ {a} {A : Set a} →
-             ≡⇒≃ (refl A) ≡ Eq.id
-  ≡⇒≃-refl = ≡⇒↝-refl
-
-  -- "Evaluation rule" for ≡⇒→.
-
-  ≡⇒→-refl : ∀ {a} {A : Set a} →
-             ≡⇒→ (refl A) ≡ id
-  ≡⇒→-refl = cong _≃_.to ≡⇒≃-refl
-
-  -- "Evaluation rule" (?) for ≃⇒≡.
-
-  ≃⇒≡-id : ∀ {a} {A : Set a}
-           (univ : Univalence′ A A) →
-           ≃⇒≡ univ Eq.id ≡ refl A
-  ≃⇒≡-id univ =
-    ≃⇒≡ univ Eq.id           ≡⟨ sym $ cong (≃⇒≡ univ) ≡⇒≃-refl ⟩
-    ≃⇒≡ univ (≡⇒≃ (refl _))  ≡⟨ _≃_.left-inverse-of (≡≃≃ univ) _ ⟩∎
-    refl _                   ∎
-
-  -- A simplification lemma for ≃⇒≡.
-
-  ≡⇒→-≃⇒≡ : ∀ k {ℓ} {A B : Set ℓ} {eq : A ≃ B}
-            (univ : Univalence′ A B) →
-            to-implication (≡⇒↝ k (≃⇒≡ univ eq)) ≡ _≃_.to eq
-  ≡⇒→-≃⇒≡ k {eq = eq} univ =
-    to-implication (≡⇒↝ k (≃⇒≡ univ eq))  ≡⟨ to-implication-≡⇒↝ k _ ⟩
-    ≡⇒↝ implication (≃⇒≡ univ eq)         ≡⟨ sym $ to-implication-≡⇒↝ equivalence _ ⟩
-    ≡⇒→ (≃⇒≡ univ eq)                     ≡⟨ cong _≃_.to (_≃_.right-inverse-of (≡≃≃ univ) _) ⟩∎
-    _≃_.to eq                             ∎
+-- "Evaluation rule" for ≡⇒≃.
+
+≡⇒≃-refl : ∀ {a} {A : Set a} →
+           ≡⇒≃ (refl A) ≡ Eq.id
+≡⇒≃-refl = ≡⇒↝-refl
+
+-- "Evaluation rule" for ≡⇒→.
+
+≡⇒→-refl : ∀ {a} {A : Set a} →
+           ≡⇒→ (refl A) ≡ id
+≡⇒→-refl = cong _≃_.to ≡⇒≃-refl
+
+-- "Evaluation rule" (?) for ≃⇒≡.
+
+≃⇒≡-id : ∀ {a} {A : Set a}
+         (univ : Univalence′ A A) →
+         ≃⇒≡ univ Eq.id ≡ refl A
+≃⇒≡-id univ =
+  ≃⇒≡ univ Eq.id           ≡⟨ sym $ cong (≃⇒≡ univ) ≡⇒≃-refl ⟩
+  ≃⇒≡ univ (≡⇒≃ (refl _))  ≡⟨ _≃_.left-inverse-of (≡≃≃ univ) _ ⟩∎
+  refl _                   ∎
+
+-- A simplification lemma for ≃⇒≡.
+
+≡⇒→-≃⇒≡ : ∀ k {ℓ} {A B : Set ℓ} {eq : A ≃ B}
+          (univ : Univalence′ A B) →
+          to-implication (≡⇒↝ k (≃⇒≡ univ eq)) ≡ _≃_.to eq
+≡⇒→-≃⇒≡ k {eq = eq} univ =
+  to-implication (≡⇒↝ k (≃⇒≡ univ eq))  ≡⟨ to-implication-≡⇒↝ k _ ⟩
+  ≡⇒↝ implication (≃⇒≡ univ eq)         ≡⟨ sym $ to-implication-≡⇒↝ equivalence _ ⟩
+  ≡⇒→ (≃⇒≡ univ eq)                     ≡⟨ cong _≃_.to (_≃_.right-inverse-of (≡≃≃ univ) _) ⟩∎
+  _≃_.to eq                             ∎
hunk ./Univalence-axiom.agda 126
-  abstract
-
-    swap≢refl : swap-as-an-equality ≢ refl Bool
-    swap≢refl =
-      swap-as-an-equality ≡ refl _            ↝⟨ cong ≡⇒→ ⟩
-      ≡⇒→ swap-as-an-equality ≡ ≡⇒→ (refl _)  ↝⟨ subst (uncurry _≡_) (cong₂ _,_ (≡⇒→-≃⇒≡ equivalence univ) ≡⇒→-refl) ⟩
-      not ≡ id                                ↝⟨ cong (_$ false) ⟩
-      true ≡ false                            ↝⟨ Bool.true≢false ⟩□
-      ⊥                                       □
-
-    -- This implies that Set is not a set.
-
-    ¬-Set-set : ¬ Is-set Set
-    ¬-Set-set =
-      Is-set Set                         ↝⟨ _⇔_.to set⇔UIP ⟩
-      Uniqueness-of-identity-proofs Set  ↝⟨ (λ uip → uip _ _) ⟩
-      swap-as-an-equality ≡ refl Bool    ↝⟨ swap≢refl ⟩□
-      ⊥                                  □
+  swap≢refl : swap-as-an-equality ≢ refl Bool
+  swap≢refl =
+    swap-as-an-equality ≡ refl _            ↝⟨ cong ≡⇒→ ⟩
+    ≡⇒→ swap-as-an-equality ≡ ≡⇒→ (refl _)  ↝⟨ subst (uncurry _≡_) (cong₂ _,_ (≡⇒→-≃⇒≡ equivalence univ) ≡⇒→-refl) ⟩
+    not ≡ id                                ↝⟨ cong (_$ false) ⟩
+    true ≡ false                            ↝⟨ Bool.true≢false ⟩□
+    ⊥                                       □
+
+  -- This implies that Set is not a set.
+
+  ¬-Set-set : ¬ Is-set Set
+  ¬-Set-set =
+    Is-set Set                         ↝⟨ _⇔_.to set⇔UIP ⟩
+    Uniqueness-of-identity-proofs Set  ↝⟨ (λ uip → uip _ _) ⟩
+    swap-as-an-equality ≡ refl Bool    ↝⟨ swap≢refl ⟩□
+    ⊥                                  □
hunk ./Univalence-axiom.agda 146
-abstract
-
-  -- Some equality types have infinitely many inhabitants (assuming
-  -- univalence).
-
-  equality-can-have-infinitely-many-inhabitants :
-    Univalence′ ℕ ℕ →
-    ∃ λ (A : Set) → ∃ λ (B : Set) →
-    ∃ λ (p : ℕ → A ≡ B) → Injective p
-  equality-can-have-infinitely-many-inhabitants univ =
-    (ℕ , ℕ , cast ∘ p , cast-preserves-injections p p-injective)
-    where
-    cast : ℕ ≃ ℕ → ℕ ≡ ℕ
-    cast = ≃⇒≡ univ
-
-    cast-preserves-injections :
-      {A : Set} (f : A → ℕ ≃ ℕ) →
-      Injective f → Injective (cast ∘ f)
-    cast-preserves-injections f inj {x = x} {y = y} cast-f-x≡cast-f-y =
-      inj (f x               ≡⟨ sym $ _≃_.right-inverse-of (≡≃≃ univ) (f x) ⟩
-           ≡⇒≃ (cast (f x))  ≡⟨ cong ≡⇒≃ cast-f-x≡cast-f-y ⟩
-           ≡⇒≃ (cast (f y))  ≡⟨ _≃_.right-inverse-of (≡≃≃ univ) (f y) ⟩∎
-           f y               ∎)
-
-    swap_and-0 : ℕ → ℕ → ℕ
-    swap i and-0 zero    = i
-    swap i and-0 (suc n) with ℕ._≟_ i (suc n)
-    ... | yes i≡1+n = zero
-    ... | no  i≢1+n = suc n
-
-    swap∘swap : ∀ i n → swap i and-0 (swap i and-0 n) ≡ n
-    swap∘swap zero    zero    = refl zero
-    swap∘swap (suc i) zero    with ℕ._≟_ i i
-    ... | yes _   = refl 0
-    ... | no  i≢i = ⊥-elim $ i≢i (refl i)
-    swap∘swap i       (suc n) with ℕ._≟_ i (suc n)
-    ... | yes i≡1+n = i≡1+n
-    ... | no  i≢1+n with ℕ._≟_ i (suc n)
-    ...   | yes i≡1+n = ⊥-elim $ i≢1+n i≡1+n
-    ...   | no  _     = refl (suc n)
-
-    p : ℕ → ℕ ≃ ℕ
-    p i = ↔⇒≃ record
-      { surjection = record
-        { logical-equivalence = record { to   = swap i and-0
-                                       ; from = swap i and-0
-                                       }
-        ; right-inverse-of    = swap∘swap i
-        }
-      ; left-inverse-of = swap∘swap i
+-- Some equality types have infinitely many inhabitants (assuming
+-- univalence).
+
+equality-can-have-infinitely-many-inhabitants :
+  Univalence′ ℕ ℕ →
+  ∃ λ (A : Set) → ∃ λ (B : Set) →
+  ∃ λ (p : ℕ → A ≡ B) → Injective p
+equality-can-have-infinitely-many-inhabitants univ =
+  (ℕ , ℕ , cast ∘ p , cast-preserves-injections p p-injective)
+  where
+  cast : ℕ ≃ ℕ → ℕ ≡ ℕ
+  cast = ≃⇒≡ univ
+
+  cast-preserves-injections :
+    {A : Set} (f : A → ℕ ≃ ℕ) →
+    Injective f → Injective (cast ∘ f)
+  cast-preserves-injections f inj {x = x} {y = y} cast-f-x≡cast-f-y =
+    inj (f x               ≡⟨ sym $ _≃_.right-inverse-of (≡≃≃ univ) (f x) ⟩
+         ≡⇒≃ (cast (f x))  ≡⟨ cong ≡⇒≃ cast-f-x≡cast-f-y ⟩
+         ≡⇒≃ (cast (f y))  ≡⟨ _≃_.right-inverse-of (≡≃≃ univ) (f y) ⟩∎
+         f y               ∎)
+
+  swap_and-0 : ℕ → ℕ → ℕ
+  swap i and-0 zero    = i
+  swap i and-0 (suc n) with ℕ._≟_ i (suc n)
+  ... | yes i≡1+n = zero
+  ... | no  i≢1+n = suc n
+
+  swap∘swap : ∀ i n → swap i and-0 (swap i and-0 n) ≡ n
+  swap∘swap zero    zero    = refl zero
+  swap∘swap (suc i) zero    with ℕ._≟_ i i
+  ... | yes _   = refl 0
+  ... | no  i≢i = ⊥-elim $ i≢i (refl i)
+  swap∘swap i       (suc n) with ℕ._≟_ i (suc n)
+  ... | yes i≡1+n = i≡1+n
+  ... | no  i≢1+n with ℕ._≟_ i (suc n)
+  ...   | yes i≡1+n = ⊥-elim $ i≢1+n i≡1+n
+  ...   | no  _     = refl (suc n)
+
+  p : ℕ → ℕ ≃ ℕ
+  p i = ↔⇒≃ record
+    { surjection = record
+      { logical-equivalence = record { to   = swap i and-0
+                                     ; from = swap i and-0
+                                     }
+      ; right-inverse-of    = swap∘swap i
hunk ./Univalence-axiom.agda 193
-
-    p-injective : Injective p
-    p-injective {x = i₁} {y = i₂} p-i₁≡p-i₂ =
-      i₁               ≡⟨ refl i₁ ⟩
-      swap i₁ and-0 0  ≡⟨ cong (λ f → f 0) swap-i₁≡swap-i₂ ⟩
-      swap i₂ and-0 0  ≡⟨ refl i₂ ⟩∎
-      i₂               ∎
-      where
-      swap-i₁≡swap-i₂ : swap i₁ and-0 ≡ swap i₂ and-0
-      swap-i₁≡swap-i₂ = cong _≃_.to p-i₁≡p-i₂
+    ; left-inverse-of = swap∘swap i
+    }
+
+  p-injective : Injective p
+  p-injective {x = i₁} {y = i₂} p-i₁≡p-i₂ =
+    i₁               ≡⟨ refl i₁ ⟩
+    swap i₁ and-0 0  ≡⟨ cong (λ f → f 0) swap-i₁≡swap-i₂ ⟩
+    swap i₂ and-0 0  ≡⟨ refl i₂ ⟩∎
+    i₂               ∎
+    where
+    swap-i₁≡swap-i₂ : swap i₁ and-0 ≡ swap i₂ and-0
+    swap-i₁≡swap-i₂ = cong _≃_.to p-i₁≡p-i₂
hunk ./Univalence-axiom.agda 209
-abstract
-
-  -- The transport theorem.
-
-  transport-theorem :
-    ∀ {p₁ p₂} (P : Set p₁ → Set p₂) →
-    (resp : ∀ {A B} → A ≃ B → P A → P B) →
-    (∀ {A} (p : P A) → resp Eq.id p ≡ p) →
-    ∀ {A B} (univ : Univalence′ A B) →
-    (A≃B : A ≃ B) (p : P A) →
-    resp A≃B p ≡ subst P (≃⇒≡ univ A≃B) p
-  transport-theorem P resp resp-id univ A≃B p =
-    resp A≃B p              ≡⟨ sym $ cong (λ q → resp q p) (right-inverse-of A≃B) ⟩
-    resp (to (from A≃B)) p  ≡⟨ elim (λ {A B} A≡B → ∀ p →
-                                       resp (≡⇒≃ A≡B) p ≡ subst P A≡B p)
-                                    (λ A p →
-                                       resp (≡⇒≃ (refl A)) p  ≡⟨ trans (cong (λ q → resp q p) ≡⇒↝-refl) (resp-id p) ⟩
-                                       p                      ≡⟨ sym $ subst-refl P p ⟩∎
-                                       subst P (refl A) p     ∎) _ _ ⟩∎
-    subst P (from A≃B) p    ∎
-    where open _≃_ (≡≃≃ univ)
-
-  -- If the univalence axiom holds, then any "resp" function that
-  -- preserves identity is an equivalence family.
-
-  resp-is-equivalence :
-    ∀ {p₁ p₂} (P : Set p₁ → Set p₂) →
-    (resp : ∀ {A B} → A ≃ B → P A → P B) →
-    (∀ {A} (p : P A) → resp Eq.id p ≡ p) →
-    ∀ {A B} (univ : Univalence′ A B) →
-    (A≃B : A ≃ B) → Is-equivalence (resp A≃B)
-  resp-is-equivalence P resp resp-id univ A≃B =
-    Eq.respects-extensional-equality
-      (λ p → sym $ transport-theorem P resp resp-id univ A≃B p)
-      (subst-is-equivalence P (≃⇒≡ univ A≃B))
-
-  -- If f is an equivalence, then (non-dependent) precomposition with
-  -- f is also an equivalence (assuming univalence).
-
-  precomposition-is-equivalence :
-    ∀ {ℓ c} {A B : Set ℓ} {C : Set c} →
-    Univalence′ B A → (A≃B : A ≃ B) →
-    Is-equivalence (λ (g : B → C) → g ∘ _≃_.to A≃B)
-  precomposition-is-equivalence {ℓ} {c} {C = C} univ A≃B =
-    resp-is-equivalence P resp refl univ (Eq.inverse A≃B)
-    where
-    P : Set ℓ → Set (ℓ ⊔ c)
-    P X = X → C
-
-    resp : ∀ {A B} → A ≃ B → P A → P B
-    resp A≃B p = p ∘ _≃_.from A≃B
+-- The transport theorem.
+
+transport-theorem :
+  ∀ {p₁ p₂} (P : Set p₁ → Set p₂) →
+  (resp : ∀ {A B} → A ≃ B → P A → P B) →
+  (∀ {A} (p : P A) → resp Eq.id p ≡ p) →
+  ∀ {A B} (univ : Univalence′ A B) →
+  (A≃B : A ≃ B) (p : P A) →
+  resp A≃B p ≡ subst P (≃⇒≡ univ A≃B) p
+transport-theorem P resp resp-id univ A≃B p =
+  resp A≃B p              ≡⟨ sym $ cong (λ q → resp q p) (right-inverse-of A≃B) ⟩
+  resp (to (from A≃B)) p  ≡⟨ elim (λ {A B} A≡B → ∀ p →
+                                     resp (≡⇒≃ A≡B) p ≡ subst P A≡B p)
+                                  (λ A p →
+                                     resp (≡⇒≃ (refl A)) p  ≡⟨ trans (cong (λ q → resp q p) ≡⇒↝-refl) (resp-id p) ⟩
+                                     p                      ≡⟨ sym $ subst-refl P p ⟩∎
+                                     subst P (refl A) p     ∎) _ _ ⟩∎
+  subst P (from A≃B) p    ∎
+  where open _≃_ (≡≃≃ univ)
+
+-- If the univalence axiom holds, then any "resp" function that
+-- preserves identity is an equivalence family.
+
+resp-is-equivalence :
+  ∀ {p₁ p₂} (P : Set p₁ → Set p₂) →
+  (resp : ∀ {A B} → A ≃ B → P A → P B) →
+  (∀ {A} (p : P A) → resp Eq.id p ≡ p) →
+  ∀ {A B} (univ : Univalence′ A B) →
+  (A≃B : A ≃ B) → Is-equivalence (resp A≃B)
+resp-is-equivalence P resp resp-id univ A≃B =
+  Eq.respects-extensional-equality
+    (λ p → sym $ transport-theorem P resp resp-id univ A≃B p)
+    (subst-is-equivalence P (≃⇒≡ univ A≃B))
+
+-- If f is an equivalence, then (non-dependent) precomposition with
+-- f is also an equivalence (assuming univalence).
+
+precomposition-is-equivalence :
+  ∀ {ℓ c} {A B : Set ℓ} {C : Set c} →
+  Univalence′ B A → (A≃B : A ≃ B) →
+  Is-equivalence (λ (g : B → C) → g ∘ _≃_.to A≃B)
+precomposition-is-equivalence {ℓ} {c} {C = C} univ A≃B =
+  resp-is-equivalence P resp refl univ (Eq.inverse A≃B)
+  where
+  P : Set ℓ → Set (ℓ ⊔ c)
+  P X = X → C
+
+  resp : ∀ {A B} → A ≃ B → P A → P B
+  resp A≃B p = p ∘ _≃_.from A≃B
hunk ./Univalence-axiom.agda 296
-abstract
-
-  -- The univalence axiom implies non-dependent functional
-  -- extensionality.
-
-  extensionality :
-    ∀ {a b} {A : Set a} {B : Set b} →
-    Univalence′ (B ²/≡) B →
-    {f g : A → B} → (∀ x → f x ≡ g x) → f ≡ g
-  extensionality {A = A} {B} univ {f} {g} f≡g =
-    f          ≡⟨ refl f ⟩
-    f′ ∘ pair  ≡⟨ cong (λ (h : B ²/≡ → B) → h ∘ pair) f′≡g′ ⟩
-    g′ ∘ pair  ≡⟨ refl g ⟩∎
-    g          ∎
-    where
-    f′ : B ²/≡ → B
-    f′ = proj₁ ∘ proj₂
-
-    g′ : B ²/≡ → B
-    g′ = proj₁
-
-    f′≡g′ : f′ ≡ g′
-    f′≡g′ = precompositions-cancel
-              univ
-              (↔⇒≃ $ Bijection.inverse -²/≡↔-)
-              (refl id)
-
-    pair : A → B ²/≡
-    pair x = (g x , f x , f≡g x)
-
-  -- The univalence axiom implies that contractibility is closed under
-  -- Π A.
-
-  Π-closure-contractible :
-    ∀ {b} → Univalence′ (Set b ²/≡) (Set b) →
-    ∀ {a} {A : Set a} {B : A → Set b} →
-    (∀ x → Univalence′ (↑ b ⊤) (B x)) →
-    (∀ x → Contractible (B x)) → Contractible ((x : A) → B x)
-  Π-closure-contractible {b} univ₁ {A = A} {B} univ₂ contr =
-    subst Contractible A→⊤≡[x:A]→Bx →⊤-contractible
-    where
-    const-⊤≡B : const (↑ b ⊤) ≡ B
-    const-⊤≡B = extensionality univ₁ λ x →
-      _≃_.from (≡≃≃ (univ₂ x)) $ ↔⇒≃ $
-        Bijection.contractible-isomorphic
-          (↑-closure 0 ⊤-contractible) (contr x)
-
-    A→⊤≡[x:A]→Bx : (A → ↑ b ⊤) ≡ ((x : A) → B x)
-    A→⊤≡[x:A]→Bx = cong (λ X → (x : A) → X x) const-⊤≡B
-
-    →⊤-contractible : Contractible (A → ↑ b ⊤)
-    →⊤-contractible = (_ , λ _ → refl _)
-
-  -- Thus we also get extensionality for dependent functions.
-
-  dependent-extensionality :
-    ∀ {b} → Univalence′ (Set b ²/≡) (Set b) →
-    ∀ {a} {A : Set a} →
-    (∀ {B : A → Set b} x → Univalence′ (↑ b ⊤) (B x)) →
-    {B : A → Set b} → Extensionality′ A B
-  dependent-extensionality univ₁ univ₂ =
-    _⇔_.to Π-closure-contractible⇔extensionality
-      (Π-closure-contractible univ₁ univ₂)
+-- The univalence axiom implies non-dependent functional
+-- extensionality.
+
+extensionality :
+  ∀ {a b} {A : Set a} {B : Set b} →
+  Univalence′ (B ²/≡) B →
+  {f g : A → B} → (∀ x → f x ≡ g x) → f ≡ g
+extensionality {A = A} {B} univ {f} {g} f≡g =
+  f          ≡⟨ refl f ⟩
+  f′ ∘ pair  ≡⟨ cong (λ (h : B ²/≡ → B) → h ∘ pair) f′≡g′ ⟩
+  g′ ∘ pair  ≡⟨ refl g ⟩∎
+  g          ∎
+  where
+  f′ : B ²/≡ → B
+  f′ = proj₁ ∘ proj₂
+
+  g′ : B ²/≡ → B
+  g′ = proj₁
+
+  f′≡g′ : f′ ≡ g′
+  f′≡g′ = precompositions-cancel
+            univ
+            (↔⇒≃ $ Bijection.inverse -²/≡↔-)
+            (refl id)
+
+  pair : A → B ²/≡
+  pair x = (g x , f x , f≡g x)
+
+-- The univalence axiom implies that contractibility is closed under
+-- Π A.
+
+Π-closure-contractible :
+  ∀ {b} → Univalence′ (Set b ²/≡) (Set b) →
+  ∀ {a} {A : Set a} {B : A → Set b} →
+  (∀ x → Univalence′ (↑ b ⊤) (B x)) →
+  (∀ x → Contractible (B x)) → Contractible ((x : A) → B x)
+Π-closure-contractible {b} univ₁ {A = A} {B} univ₂ contr =
+  subst Contractible A→⊤≡[x:A]→Bx →⊤-contractible
+  where
+  const-⊤≡B : const (↑ b ⊤) ≡ B
+  const-⊤≡B = extensionality univ₁ λ x →
+    _≃_.from (≡≃≃ (univ₂ x)) $ ↔⇒≃ $
+      Bijection.contractible-isomorphic
+        (↑-closure 0 ⊤-contractible) (contr x)
+
+  A→⊤≡[x:A]→Bx : (A → ↑ b ⊤) ≡ ((x : A) → B x)
+  A→⊤≡[x:A]→Bx = cong (λ X → (x : A) → X x) const-⊤≡B
+
+  →⊤-contractible : Contractible (A → ↑ b ⊤)
+  →⊤-contractible = (_ , λ _ → refl _)
+
+-- Thus we also get extensionality for dependent functions.
+
+dependent-extensionality :
+  ∀ {b} → Univalence′ (Set b ²/≡) (Set b) →
+  ∀ {a} {A : Set a} →
+  (∀ {B : A → Set b} x → Univalence′ (↑ b ⊤) (B x)) →
+  {B : A → Set b} → Extensionality′ A B
+dependent-extensionality univ₁ univ₂ =
+  _⇔_.to Π-closure-contractible⇔extensionality
+    (Π-closure-contractible univ₁ univ₂)
hunk ./Univalence-axiom.agda 361
-abstract
-
-  -- The univalence axiom is propositional (assuming extensionality).
-
-  Univalence′-propositional :
-    ∀ {ℓ} → Extensionality (lsuc ℓ) (lsuc ℓ) →
-    {A B : Set ℓ} → Is-proposition (Univalence′ A B)
-  Univalence′-propositional ext =
-    Eq.propositional ext ≡⇒≃
-
-  Univalence-propositional :
-    ∀ {ℓ} → Extensionality (lsuc ℓ) (lsuc ℓ) →
-    Is-proposition (Univalence ℓ)
-  Univalence-propositional ext =
-    implicit-Π-closure ext 1 λ _ →
-    implicit-Π-closure ext 1 λ _ →
-    Univalence′-propositional ext
-
-  -- ≡⇒≃ commutes with sym/Eq.inverse (assuming extensionality).
-
-  ≡⇒≃-sym :
-    ∀ {ℓ} → Extensionality ℓ ℓ →
-    {A B : Set ℓ} (A≡B : A ≡ B) →
-    ≡⇒≃ (sym A≡B) ≡ Eq.inverse (≡⇒≃ A≡B)
-  ≡⇒≃-sym ext = elim¹
-
-    (λ eq → ≡⇒≃ (sym eq) ≡ Eq.inverse (≡⇒≃ eq))
-
-    (≡⇒≃ (sym (refl _))         ≡⟨ cong ≡⇒≃ sym-refl ⟩
-     ≡⇒≃ (refl _)               ≡⟨ ≡⇒≃-refl ⟩
-     Eq.id                      ≡⟨ sym $ Groupoid.identity (Eq.groupoid ext) ⟩
-     Eq.inverse Eq.id           ≡⟨ cong Eq.inverse $ sym ≡⇒≃-refl ⟩∎
-     Eq.inverse (≡⇒≃ (refl _))  ∎)
-
-  -- ≃⇒≡ univ commutes with Eq.inverse/sym (assuming extensionality).
-
-  ≃⇒≡-inverse :
-    ∀ {ℓ} (univ : Univalence ℓ) → Extensionality ℓ ℓ →
-    {A B : Set ℓ} (A≃B : A ≃ B) →
-    ≃⇒≡ univ (Eq.inverse A≃B) ≡ sym (≃⇒≡ univ A≃B)
-  ≃⇒≡-inverse univ ext A≃B =
-    ≃⇒≡ univ (Eq.inverse A≃B)                   ≡⟨ sym $ cong (λ p → ≃⇒≡ univ (Eq.inverse p)) (_≃_.right-inverse-of (≡≃≃ univ) _) ⟩
-    ≃⇒≡ univ (Eq.inverse (≡⇒≃ (≃⇒≡ univ A≃B)))  ≡⟨ cong (≃⇒≡ univ) $ sym $ ≡⇒≃-sym ext _ ⟩
-    ≃⇒≡ univ (≡⇒≃ (sym (≃⇒≡ univ A≃B)))         ≡⟨ _≃_.left-inverse-of (≡≃≃ univ) _ ⟩∎
-    sym (≃⇒≡ univ A≃B)                          ∎
-
-  -- ≡⇒≃ commutes with trans/_⊚_ (assuming extensionality).
-
-  ≡⇒≃-trans :
-    ∀ {ℓ} → Extensionality ℓ ℓ →
-    {A B C : Set ℓ} (A≡B : A ≡ B) (B≡C : B ≡ C) →
-    ≡⇒≃ (trans A≡B B≡C) ≡ ≡⇒≃ B≡C ⊚ ≡⇒≃ A≡B
-  ≡⇒≃-trans ext A≡B = elim¹
-
-    (λ eq → ≡⇒≃ (trans A≡B eq) ≡ ≡⇒≃ eq ⊚ ≡⇒≃ A≡B)
-
-    (≡⇒≃ (trans A≡B (refl _))  ≡⟨ cong ≡⇒≃ $ trans-reflʳ _ ⟩
-     ≡⇒≃ A≡B                   ≡⟨ sym $ Groupoid.left-identity (Eq.groupoid ext) _ ⟩
-     Eq.id ⊚ ≡⇒≃ A≡B           ≡⟨ sym $ cong (λ eq → eq ⊚ ≡⇒≃ A≡B) ≡⇒≃-refl ⟩∎
-     ≡⇒≃ (refl _) ⊚ ≡⇒≃ A≡B    ∎)
-
-  -- ≃⇒≡ univ commutes with _⊚_/flip trans (assuming extensionality).
-
-  ≃⇒≡-∘ :
-    ∀ {ℓ} (univ : Univalence ℓ) → Extensionality ℓ ℓ →
-    {A B C : Set ℓ} (A≃B : A ≃ B) (B≃C : B ≃ C) →
-    ≃⇒≡ univ (B≃C ⊚ A≃B) ≡ trans (≃⇒≡ univ A≃B) (≃⇒≡ univ B≃C)
-  ≃⇒≡-∘ univ ext A≃B B≃C =
-    ≃⇒≡ univ (B≃C ⊚ A≃B)                                  ≡⟨ sym $ cong₂ (λ p q → ≃⇒≡ univ (p ⊚ q)) (_≃_.right-inverse-of (≡≃≃ univ) _)
-                                                                                                    (_≃_.right-inverse-of (≡≃≃ univ) _) ⟩
-    ≃⇒≡ univ (≡⇒≃ (≃⇒≡ univ B≃C) ⊚ ≡⇒≃ (≃⇒≡ univ A≃B))    ≡⟨ cong (≃⇒≡ univ) $ sym $ ≡⇒≃-trans ext _ _ ⟩
-    ≃⇒≡ univ (≡⇒≃ (trans (≃⇒≡ univ A≃B) (≃⇒≡ univ B≃C)))  ≡⟨ _≃_.left-inverse-of (≡≃≃ univ) _ ⟩∎
-    trans (≃⇒≡ univ A≃B) (≃⇒≡ univ B≃C)                   ∎
-
-  -- A variant of the transport theorem.
-
-  transport-theorem′ :
-    ∀ {a p r} {A : Set a}
-    (P : A → Set p) (R : A → A → Set r)
-    (≡↠R : ∀ {x y} → (x ≡ y) ↠ R x y)
-    (resp : ∀ {x y} → R x y → P x → P y) →
-    (∀ x p → resp (_↠_.to ≡↠R (refl x)) p ≡ p) →
-    ∀ {x y} (r : R x y) (p : P x) →
-    resp r p ≡ subst P (_↠_.from ≡↠R r) p
-  transport-theorem′ P R ≡↠R resp hyp r p =
-    resp r p              ≡⟨ sym $ cong (λ r → resp r p) (right-inverse-of r) ⟩
-    resp (to (from r)) p  ≡⟨ elim (λ {x y} x≡y → ∀ p →
-                                     resp (_↠_.to ≡↠R x≡y) p ≡ subst P x≡y p)
-                                  (λ x p →
-                                     resp (_↠_.to ≡↠R (refl x)) p  ≡⟨ hyp x p ⟩
-                                     p                             ≡⟨ sym $ subst-refl P p ⟩∎
-                                     subst P (refl x) p            ∎) _ _ ⟩∎
-    subst P (from r) p    ∎
-    where open _↠_ ≡↠R
-
-  -- Simplification (?) lemma for transport-theorem′.
-
-  transport-theorem′-refl :
-    ∀ {a p r} {A : Set a}
-    (P : A → Set p) (R : A → A → Set r)
-    (≡≃R : ∀ {x y} → (x ≡ y) ≃ R x y)
-    (resp : ∀ {x y} → R x y → P x → P y) →
-    (resp-refl : ∀ x p → resp (_≃_.to ≡≃R (refl x)) p ≡ p) →
-    ∀ {x} (p : P x) →
-    transport-theorem′ P R (_≃_.surjection ≡≃R) resp resp-refl
-                       (_≃_.to ≡≃R (refl x)) p ≡
-    trans (trans (resp-refl x p) (sym $ subst-refl P p))
-          (sym $ cong (λ eq → subst P eq p)
-                      (_≃_.left-inverse-of ≡≃R (refl x)))
-  transport-theorem′-refl P R ≡≃R resp resp-refl {x} p =
-
-    let body = λ x p → trans (resp-refl x p) (sym $ subst-refl P p)
-
-        lemma =
-          trans (sym $ cong (λ r → resp (to r) p) $ refl (refl x))
-                (elim (λ {x y} x≡y →
-                         ∀ p → resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
-                      (λ x p → trans (resp-refl x p)
-                                     (sym $ subst-refl P p))
-                      (refl x) p)                                        ≡⟨ cong₂ (λ q r → trans q (r p))
-                                                                                  (sym $ cong-sym (λ r → resp (to r) p) _)
-                                                                                  (elim-refl (λ {x y} x≡y → ∀ p →
-                                                                                                resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p) _) ⟩
-          trans (cong (λ r → resp (to r) p) $ sym $ refl (refl x))
-                (body x p)                                               ≡⟨ cong (λ q → trans (cong (λ r → resp (to r) p) q) (body x p))
-                                                                                 sym-refl ⟩
-          trans (cong (λ r → resp (to r) p) $ refl (refl x)) (body x p)  ≡⟨ cong (λ q → trans q (body x p)) $
-                                                                                 cong-refl (λ r → resp (to r) p) ⟩
-          trans (refl (resp (to (refl x)) p)) (body x p)                 ≡⟨ trans-reflˡ _ ⟩
-
-          body x p                                                       ≡⟨ sym $ trans-reflʳ _ ⟩
-
-          trans (body x p) (refl (subst P (refl x) p))                   ≡⟨ cong (trans (body x p)) $ sym $
-                                                                                 cong-refl (λ eq → subst P eq p) ⟩
-          trans (body x p)
-                (cong (λ eq → subst P eq p) (refl (refl x)))             ≡⟨ cong (trans (body x p) ∘ cong (λ eq → subst P eq p)) $
-                                                                                 sym sym-refl ⟩
-          trans (body x p)
-                (cong (λ eq → subst P eq p) (sym (refl (refl x))))       ≡⟨ cong (trans (body x p)) $ cong-sym (λ eq → subst P eq p) _ ⟩∎
-
-          trans (body x p)
-                (sym $ cong (λ eq → subst P eq p) (refl (refl x)))       ∎
-    in
-
-    trans (sym $ cong (λ r → resp r p) $ right-inverse-of (to (refl x)))
-          (elim (λ {x y} x≡y →
-                   ∀ p → resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
-                body (from (to (refl x))) p)                              ≡⟨ cong (λ eq → trans (sym $ cong (λ r → resp r p) eq)
-                                                                                                (elim (λ {x y} x≡y → ∀ p →
-                                                                                                         resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
-                                                                                                      body (from (to (refl x))) p)) $
-                                                                                  sym $ left-right-lemma (refl x) ⟩
-    trans (sym $ cong (λ r → resp r p) $ cong to $
-             left-inverse-of (refl x))
-          (elim (λ {x y} x≡y →
-                   ∀ p → resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
-                body (from (to (refl x))) p)                              ≡⟨ cong (λ eq → trans (sym eq)
-                                                                                                (elim (λ {x y} x≡y → ∀ p →
-                                                                                                         resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
-                                                                                                      body (from (to (refl x))) p)) $
-                                                                                  cong-∘ (λ r → resp r p) to _ ⟩
-    trans (sym $ cong (λ r → resp (to r) p) $ left-inverse-of (refl x))
-          (elim (λ {x y} x≡y →
-                   ∀ p → resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
-                body (from (to (refl x))) p)                              ≡⟨ elim₁ (λ {q} eq → trans (sym $ cong (λ r → resp (to r) p) eq)
-                                                                                                     (elim (λ {x y} x≡y → ∀ p →
-                                                                                                              resp (_≃_.to ≡≃R x≡y) p ≡
-                                                                                                              subst P x≡y p)
-                                                                                                           body q p) ≡
-                                                                                               trans (body x p)
-                                                                                                     (sym $ cong (λ eq → subst P eq p) eq))
-                                                                                   lemma
-                                                                                   (left-inverse-of (refl x)) ⟩∎
-
-    trans (body x p)
-          (sym $ cong (λ eq → subst P eq p) (left-inverse-of (refl x)))   ∎
-
-    where open _≃_ ≡≃R
-
-  -- Simplification (?) lemma for transport-theorem.
-
-  transport-theorem-≡⇒≃-refl :
-    ∀ {p₁ p₂} (P : Set p₁ → Set p₂)
-    (resp : ∀ {A B} → A ≃ B → P A → P B)
-    (resp-id : ∀ {A} (p : P A) → resp Eq.id p ≡ p)
-    (univ : Univalence p₁) {A} (p : P A) →
-    transport-theorem P resp resp-id univ (≡⇒≃ (refl A)) p ≡
-    trans (trans (trans (cong (λ eq → resp eq p) ≡⇒≃-refl)
-                    (resp-id p))
-             (sym $ subst-refl P p))
-      (sym $ cong (λ eq → subst P eq p)
-                  (_≃_.left-inverse-of (≡≃≃ univ) (refl A)))
-  transport-theorem-≡⇒≃-refl P resp resp-id univ {A} p =
-    transport-theorem′-refl P _≃_ (≡≃≃ univ) resp
-      (λ _ p → trans (cong (λ eq → resp eq p) ≡⇒≃-refl) (resp-id p)) p
-
-  -- A variant of resp-is-equivalence.
-
-  resp-is-equivalence′ :
-    ∀ {a p r} {A : Set a}
-    (P : A → Set p) (R : A → A → Set r)
-    (≡↠R : ∀ {x y} → (x ≡ y) ↠ R x y)
-    (resp : ∀ {x y} → R x y → P x → P y) →
-    (∀ x p → resp (_↠_.to ≡↠R (refl x)) p ≡ p) →
-    ∀ {x y} (r : R x y) → Is-equivalence (resp r)
-  resp-is-equivalence′ P R ≡↠R resp hyp r =
-    Eq.respects-extensional-equality
-      (λ p → sym $ transport-theorem′ P R ≡↠R resp hyp r p)
-      (subst-is-equivalence P (_↠_.from ≡↠R r))
-
-  -- A lemma relating ≃⇒≡, →-cong and cong₂.
-
-  ≃⇒≡-→-cong :
-    ∀ {ℓ} {A₁ A₂ B₁ B₂ : Set ℓ}
-    (ext : Extensionality ℓ ℓ) →
-    (univ : Univalence ℓ)
-    (A₁≃A₂ : A₁ ≃ A₂) (B₁≃B₂ : B₁ ≃ B₂) →
-    ≃⇒≡ univ (→-cong ext A₁≃A₂ B₁≃B₂) ≡
-      cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂) (≃⇒≡ univ B₁≃B₂)
-  ≃⇒≡-→-cong {A₂ = A₂} {B₁} ext univ A₁≃A₂ B₁≃B₂ =
-    ≃⇒≡ univ (→-cong ext A₁≃A₂ B₁≃B₂)                        ≡⟨ cong (≃⇒≡ univ) (lift-equality ext lemma) ⟩
-
-    ≃⇒≡ univ (≡⇒≃ (cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂)
-                                         (≃⇒≡ univ B₁≃B₂)))  ≡⟨ left-inverse-of (≡≃≃ univ) _ ⟩∎
-
-    cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂) (≃⇒≡ univ B₁≃B₂)  ∎
-    where
-    open _≃_
-
-    lemma :
-      (λ f → to B₁≃B₂ ∘ f ∘ from A₁≃A₂) ≡
-      to (≡⇒≃ (cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂)
-                                     (≃⇒≡ univ B₁≃B₂)))
-    lemma =
-      (λ f → to B₁≃B₂ ∘ f ∘ from A₁≃A₂)                  ≡⟨ ext (λ _ → transport-theorem (λ B → A₂ → B) (λ A≃B g → _≃_.to A≃B ∘ g)
-                                                                                         refl univ B₁≃B₂ _) ⟩
-      subst (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂) ∘
-      (λ f → f ∘ from A₁≃A₂)                             ≡⟨ cong (_∘_ (subst (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂))) (ext λ f →
-                                                              transport-theorem (λ A → A → B₁) (λ A≃B g → g ∘ _≃_.from A≃B) refl univ A₁≃A₂ f) ⟩
-      subst (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂) ∘
-      subst (λ A → A → B₁) (≃⇒≡ univ A₁≃A₂)              ≡⟨ cong₂ (λ g h f → g (h f)) (ext $ subst-in-terms-of-≡⇒↝ equivalence _ (λ B → A₂ → B))
-                                                                                      (ext $ subst-in-terms-of-≡⇒↝ equivalence _ (λ A → A → B₁)) ⟩
-      to (≡⇒≃ (cong (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂))) ∘
-      to (≡⇒≃ (cong (λ A → A → B₁) (≃⇒≡ univ A₁≃A₂)))    ≡⟨⟩
-
-      to (≡⇒≃ (cong (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂)) ⊚
-          ≡⇒≃ (cong (λ A → A → B₁) (≃⇒≡ univ A₁≃A₂)))    ≡⟨ cong to $ sym $ ≡⇒≃-trans ext _ _ ⟩∎
-
-      to (≡⇒≃ (cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂)
-                                     (≃⇒≡ univ B₁≃B₂)))  ∎
-
-  -- One can sometimes push cong through ≃⇒≡ (assuming
-  -- extensionality).
-
-  cong-≃⇒≡ :
-    ∀ {ℓ p} {A B : Set ℓ} {A≃B : A ≃ B} →
-    Extensionality p p →
-    (univ₁ : Univalence ℓ)
-    (univ₂ : Univalence p)
-    (P : Set ℓ → Set p)
-    (P-cong : ∀ {A B} → A ≃ B → P A ≃ P B) →
-    (∀ {A} (p : P A) → _≃_.to (P-cong Eq.id) p ≡ p) →
-    cong P (≃⇒≡ univ₁ A≃B) ≡ ≃⇒≡ univ₂ (P-cong A≃B)
-  cong-≃⇒≡ {A≃B = A≃B} ext univ₁ univ₂ P P-cong P-cong-id =
-    cong P (≃⇒≡ univ₁ A≃B)                    ≡⟨ sym $ _≃_.left-inverse-of (≡≃≃ univ₂) _ ⟩
-    ≃⇒≡ univ₂ (≡⇒≃ (cong P (≃⇒≡ univ₁ A≃B)))  ≡⟨ cong (≃⇒≡ univ₂) $ lift-equality ext lemma ⟩∎
-    ≃⇒≡ univ₂ (P-cong A≃B)                    ∎
-    where
-    lemma : ≡⇒→ (cong P (≃⇒≡ univ₁ A≃B)) ≡ _≃_.to (P-cong A≃B)
-    lemma = ext λ x →
-      ≡⇒→ (cong P (≃⇒≡ univ₁ A≃B)) x  ≡⟨ sym $ subst-in-terms-of-≡⇒↝ equivalence _ P x ⟩
-      subst P (≃⇒≡ univ₁ A≃B) x       ≡⟨ sym $ transport-theorem P (_≃_.to ∘ P-cong) P-cong-id univ₁ A≃B x ⟩∎
-      _≃_.to (P-cong A≃B) x           ∎
-
-  -- Any "resp" function that preserves identity also preserves
-  -- compositions (assuming univalence and extensionality).
-
-  resp-preserves-compositions :
-    ∀ {p₁ p₂} (P : Set p₁ → Set p₂) →
-    (resp : ∀ {A B} → A ≃ B → P A → P B) →
-    (∀ {A} (p : P A) → resp Eq.id p ≡ p) →
-    Univalence p₁ → Extensionality p₁ p₁ →
-    ∀ {A B C} (A≃B : A ≃ B) (B≃C : B ≃ C) p →
-    resp (B≃C ⊚ A≃B) p ≡ (resp B≃C ∘ resp A≃B) p
-  resp-preserves-compositions P resp resp-id univ ext A≃B B≃C p =
-    resp (B≃C ⊚ A≃B) p                                 ≡⟨ transport-theorem P resp resp-id univ _ _ ⟩
-    subst P (≃⇒≡ univ (B≃C ⊚ A≃B)) p                   ≡⟨ cong (λ eq → subst P eq p) $ ≃⇒≡-∘ univ ext A≃B B≃C ⟩
-    subst P (trans (≃⇒≡ univ A≃B) (≃⇒≡ univ B≃C)) p    ≡⟨ sym $ subst-subst P _ _ _ ⟩
-    subst P (≃⇒≡ univ B≃C) (subst P (≃⇒≡ univ A≃B) p)  ≡⟨ sym $ transport-theorem P resp resp-id univ _ _ ⟩
-    resp B≃C (subst P (≃⇒≡ univ A≃B) p)                ≡⟨ sym $ cong (resp _) $ transport-theorem P resp resp-id univ _ _ ⟩∎
-    resp B≃C (resp A≃B p)                              ∎
-
-  -- Any "resp" function that preserves identity also preserves
-  -- inverses (assuming univalence and extensionality).
-
-  resp-preserves-inverses :
-    ∀ {p₁ p₂} (P : Set p₁ → Set p₂) →
-    (resp : ∀ {A B} → A ≃ B → P A → P B) →
-    (∀ {A} (p : P A) → resp Eq.id p ≡ p) →
-    Univalence p₁ → Extensionality p₁ p₁ →
-    ∀ {A B} (A≃B : A ≃ B) p q →
-    resp A≃B p ≡ q → resp (inverse A≃B) q ≡ p
-  resp-preserves-inverses P resp resp-id univ ext A≃B p q eq =
-    let lemma =
-          q                                     ≡⟨ sym eq ⟩
-          resp A≃B p                            ≡⟨ transport-theorem P resp resp-id univ _ _ ⟩
-          subst P (≃⇒≡ univ A≃B) p              ≡⟨ cong (λ eq → subst P eq p) $ sym $ sym-sym _ ⟩∎
-          subst P (sym (sym (≃⇒≡ univ A≃B))) p  ∎
-    in
-
-    resp (inverse A≃B) q                                                 ≡⟨ transport-theorem P resp resp-id univ _ _ ⟩
-    subst P (≃⇒≡ univ (inverse A≃B)) q                                   ≡⟨ cong (λ eq → subst P eq q) $ ≃⇒≡-inverse univ ext A≃B ⟩
-    subst P (sym (≃⇒≡ univ A≃B)) q                                       ≡⟨ cong (subst P (sym (≃⇒≡ univ A≃B))) lemma ⟩
-    subst P (sym (≃⇒≡ univ A≃B)) (subst P (sym (sym (≃⇒≡ univ A≃B))) p)  ≡⟨ subst-subst-sym P _ _ ⟩∎
-    p                                                                    ∎
+-- The univalence axiom is propositional (assuming extensionality).
+
+Univalence′-propositional :
+  ∀ {ℓ} → Extensionality (lsuc ℓ) (lsuc ℓ) →
+  {A B : Set ℓ} → Is-proposition (Univalence′ A B)
+Univalence′-propositional ext =
+  Eq.propositional ext ≡⇒≃
+
+Univalence-propositional :
+  ∀ {ℓ} → Extensionality (lsuc ℓ) (lsuc ℓ) →
+  Is-proposition (Univalence ℓ)
+Univalence-propositional ext =
+  implicit-Π-closure ext 1 λ _ →
+  implicit-Π-closure ext 1 λ _ →
+  Univalence′-propositional ext
+
+-- ≡⇒≃ commutes with sym/Eq.inverse (assuming extensionality).
+
+≡⇒≃-sym :
+  ∀ {ℓ} → Extensionality ℓ ℓ →
+  {A B : Set ℓ} (A≡B : A ≡ B) →
+  ≡⇒≃ (sym A≡B) ≡ Eq.inverse (≡⇒≃ A≡B)
+≡⇒≃-sym ext = elim¹
+
+  (λ eq → ≡⇒≃ (sym eq) ≡ Eq.inverse (≡⇒≃ eq))
+
+  (≡⇒≃ (sym (refl _))         ≡⟨ cong ≡⇒≃ sym-refl ⟩
+   ≡⇒≃ (refl _)               ≡⟨ ≡⇒≃-refl ⟩
+   Eq.id                      ≡⟨ sym $ Groupoid.identity (Eq.groupoid ext) ⟩
+   Eq.inverse Eq.id           ≡⟨ cong Eq.inverse $ sym ≡⇒≃-refl ⟩∎
+   Eq.inverse (≡⇒≃ (refl _))  ∎)
+
+-- ≃⇒≡ univ commutes with Eq.inverse/sym (assuming extensionality).
+
+≃⇒≡-inverse :
+  ∀ {ℓ} (univ : Univalence ℓ) → Extensionality ℓ ℓ →
+  {A B : Set ℓ} (A≃B : A ≃ B) →
+  ≃⇒≡ univ (Eq.inverse A≃B) ≡ sym (≃⇒≡ univ A≃B)
+≃⇒≡-inverse univ ext A≃B =
+  ≃⇒≡ univ (Eq.inverse A≃B)                   ≡⟨ sym $ cong (λ p → ≃⇒≡ univ (Eq.inverse p)) (_≃_.right-inverse-of (≡≃≃ univ) _) ⟩
+  ≃⇒≡ univ (Eq.inverse (≡⇒≃ (≃⇒≡ univ A≃B)))  ≡⟨ cong (≃⇒≡ univ) $ sym $ ≡⇒≃-sym ext _ ⟩
+  ≃⇒≡ univ (≡⇒≃ (sym (≃⇒≡ univ A≃B)))         ≡⟨ _≃_.left-inverse-of (≡≃≃ univ) _ ⟩∎
+  sym (≃⇒≡ univ A≃B)                          ∎
+
+-- ≡⇒≃ commutes with trans/_⊚_ (assuming extensionality).
+
+≡⇒≃-trans :
+  ∀ {ℓ} → Extensionality ℓ ℓ →
+  {A B C : Set ℓ} (A≡B : A ≡ B) (B≡C : B ≡ C) →
+  ≡⇒≃ (trans A≡B B≡C) ≡ ≡⇒≃ B≡C ⊚ ≡⇒≃ A≡B
+≡⇒≃-trans ext A≡B = elim¹
+
+  (λ eq → ≡⇒≃ (trans A≡B eq) ≡ ≡⇒≃ eq ⊚ ≡⇒≃ A≡B)
+
+  (≡⇒≃ (trans A≡B (refl _))  ≡⟨ cong ≡⇒≃ $ trans-reflʳ _ ⟩
+   ≡⇒≃ A≡B                   ≡⟨ sym $ Groupoid.left-identity (Eq.groupoid ext) _ ⟩
+   Eq.id ⊚ ≡⇒≃ A≡B           ≡⟨ sym $ cong (λ eq → eq ⊚ ≡⇒≃ A≡B) ≡⇒≃-refl ⟩∎
+   ≡⇒≃ (refl _) ⊚ ≡⇒≃ A≡B    ∎)
+
+-- ≃⇒≡ univ commutes with _⊚_/flip trans (assuming extensionality).
+
+≃⇒≡-∘ :
+  ∀ {ℓ} (univ : Univalence ℓ) → Extensionality ℓ ℓ →
+  {A B C : Set ℓ} (A≃B : A ≃ B) (B≃C : B ≃ C) →
+  ≃⇒≡ univ (B≃C ⊚ A≃B) ≡ trans (≃⇒≡ univ A≃B) (≃⇒≡ univ B≃C)
+≃⇒≡-∘ univ ext A≃B B≃C =
+  ≃⇒≡ univ (B≃C ⊚ A≃B)                                  ≡⟨ sym $ cong₂ (λ p q → ≃⇒≡ univ (p ⊚ q)) (_≃_.right-inverse-of (≡≃≃ univ) _)
+                                                                                                  (_≃_.right-inverse-of (≡≃≃ univ) _) ⟩
+  ≃⇒≡ univ (≡⇒≃ (≃⇒≡ univ B≃C) ⊚ ≡⇒≃ (≃⇒≡ univ A≃B))    ≡⟨ cong (≃⇒≡ univ) $ sym $ ≡⇒≃-trans ext _ _ ⟩
+  ≃⇒≡ univ (≡⇒≃ (trans (≃⇒≡ univ A≃B) (≃⇒≡ univ B≃C)))  ≡⟨ _≃_.left-inverse-of (≡≃≃ univ) _ ⟩∎
+  trans (≃⇒≡ univ A≃B) (≃⇒≡ univ B≃C)                   ∎
+
+-- A variant of the transport theorem.
+
+transport-theorem′ :
+  ∀ {a p r} {A : Set a}
+  (P : A → Set p) (R : A → A → Set r)
+  (≡↠R : ∀ {x y} → (x ≡ y) ↠ R x y)
+  (resp : ∀ {x y} → R x y → P x → P y) →
+  (∀ x p → resp (_↠_.to ≡↠R (refl x)) p ≡ p) →
+  ∀ {x y} (r : R x y) (p : P x) →
+  resp r p ≡ subst P (_↠_.from ≡↠R r) p
+transport-theorem′ P R ≡↠R resp hyp r p =
+  resp r p              ≡⟨ sym $ cong (λ r → resp r p) (right-inverse-of r) ⟩
+  resp (to (from r)) p  ≡⟨ elim (λ {x y} x≡y → ∀ p →
+                                   resp (_↠_.to ≡↠R x≡y) p ≡ subst P x≡y p)
+                                (λ x p →
+                                   resp (_↠_.to ≡↠R (refl x)) p  ≡⟨ hyp x p ⟩
+                                   p                             ≡⟨ sym $ subst-refl P p ⟩∎
+                                   subst P (refl x) p            ∎) _ _ ⟩∎
+  subst P (from r) p    ∎
+  where open _↠_ ≡↠R
+
+-- Simplification (?) lemma for transport-theorem′.
+
+transport-theorem′-refl :
+  ∀ {a p r} {A : Set a}
+  (P : A → Set p) (R : A → A → Set r)
+  (≡≃R : ∀ {x y} → (x ≡ y) ≃ R x y)
+  (resp : ∀ {x y} → R x y → P x → P y) →
+  (resp-refl : ∀ x p → resp (_≃_.to ≡≃R (refl x)) p ≡ p) →
+  ∀ {x} (p : P x) →
+  transport-theorem′ P R (_≃_.surjection ≡≃R) resp resp-refl
+                     (_≃_.to ≡≃R (refl x)) p ≡
+  trans (trans (resp-refl x p) (sym $ subst-refl P p))
+        (sym $ cong (λ eq → subst P eq p)
+                    (_≃_.left-inverse-of ≡≃R (refl x)))
+transport-theorem′-refl P R ≡≃R resp resp-refl {x} p =
+
+  let body = λ x p → trans (resp-refl x p) (sym $ subst-refl P p)
+
+      lemma =
+        trans (sym $ cong (λ r → resp (to r) p) $ refl (refl x))
+              (elim (λ {x y} x≡y →
+                       ∀ p → resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
+                    (λ x p → trans (resp-refl x p)
+                                   (sym $ subst-refl P p))
+                    (refl x) p)                                        ≡⟨ cong₂ (λ q r → trans q (r p))
+                                                                                (sym $ cong-sym (λ r → resp (to r) p) _)
+                                                                                (elim-refl (λ {x y} x≡y → ∀ p →
+                                                                                              resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p) _) ⟩
+        trans (cong (λ r → resp (to r) p) $ sym $ refl (refl x))
+              (body x p)                                               ≡⟨ cong (λ q → trans (cong (λ r → resp (to r) p) q) (body x p))
+                                                                               sym-refl ⟩
+        trans (cong (λ r → resp (to r) p) $ refl (refl x)) (body x p)  ≡⟨ cong (λ q → trans q (body x p)) $
+                                                                               cong-refl (λ r → resp (to r) p) ⟩
+        trans (refl (resp (to (refl x)) p)) (body x p)                 ≡⟨ trans-reflˡ _ ⟩
+
+        body x p                                                       ≡⟨ sym $ trans-reflʳ _ ⟩
+
+        trans (body x p) (refl (subst P (refl x) p))                   ≡⟨ cong (trans (body x p)) $ sym $
+                                                                               cong-refl (λ eq → subst P eq p) ⟩
+        trans (body x p)
+              (cong (λ eq → subst P eq p) (refl (refl x)))             ≡⟨ cong (trans (body x p) ∘ cong (λ eq → subst P eq p)) $
+                                                                               sym sym-refl ⟩
+        trans (body x p)
+              (cong (λ eq → subst P eq p) (sym (refl (refl x))))       ≡⟨ cong (trans (body x p)) $ cong-sym (λ eq → subst P eq p) _ ⟩∎
+
+        trans (body x p)
+              (sym $ cong (λ eq → subst P eq p) (refl (refl x)))       ∎
+  in
+
+  trans (sym $ cong (λ r → resp r p) $ right-inverse-of (to (refl x)))
+        (elim (λ {x y} x≡y →
+                 ∀ p → resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
+              body (from (to (refl x))) p)                              ≡⟨ cong (λ eq → trans (sym $ cong (λ r → resp r p) eq)
+                                                                                              (elim (λ {x y} x≡y → ∀ p →
+                                                                                                       resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
+                                                                                                    body (from (to (refl x))) p)) $
+                                                                                sym $ left-right-lemma (refl x) ⟩
+  trans (sym $ cong (λ r → resp r p) $ cong to $
+           left-inverse-of (refl x))
+        (elim (λ {x y} x≡y →
+                 ∀ p → resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
+              body (from (to (refl x))) p)                              ≡⟨ cong (λ eq → trans (sym eq)
+                                                                                              (elim (λ {x y} x≡y → ∀ p →
+                                                                                                       resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
+                                                                                                    body (from (to (refl x))) p)) $
+                                                                                cong-∘ (λ r → resp r p) to _ ⟩
+  trans (sym $ cong (λ r → resp (to r) p) $ left-inverse-of (refl x))
+        (elim (λ {x y} x≡y →
+                 ∀ p → resp (_≃_.to ≡≃R x≡y) p ≡ subst P x≡y p)
+              body (from (to (refl x))) p)                              ≡⟨ elim₁ (λ {q} eq → trans (sym $ cong (λ r → resp (to r) p) eq)
+                                                                                                   (elim (λ {x y} x≡y → ∀ p →
+                                                                                                            resp (_≃_.to ≡≃R x≡y) p ≡
+                                                                                                            subst P x≡y p)
+                                                                                                         body q p) ≡
+                                                                                             trans (body x p)
+                                                                                                   (sym $ cong (λ eq → subst P eq p) eq))
+                                                                                 lemma
+                                                                                 (left-inverse-of (refl x)) ⟩∎
+
+  trans (body x p)
+        (sym $ cong (λ eq → subst P eq p) (left-inverse-of (refl x)))   ∎
+
+  where open _≃_ ≡≃R
+
+-- Simplification (?) lemma for transport-theorem.
+
+transport-theorem-≡⇒≃-refl :
+  ∀ {p₁ p₂} (P : Set p₁ → Set p₂)
+  (resp : ∀ {A B} → A ≃ B → P A → P B)
+  (resp-id : ∀ {A} (p : P A) → resp Eq.id p ≡ p)
+  (univ : Univalence p₁) {A} (p : P A) →
+  transport-theorem P resp resp-id univ (≡⇒≃ (refl A)) p ≡
+  trans (trans (trans (cong (λ eq → resp eq p) ≡⇒≃-refl)
+                  (resp-id p))
+           (sym $ subst-refl P p))
+    (sym $ cong (λ eq → subst P eq p)
+                (_≃_.left-inverse-of (≡≃≃ univ) (refl A)))
+transport-theorem-≡⇒≃-refl P resp resp-id univ {A} p =
+  transport-theorem′-refl P _≃_ (≡≃≃ univ) resp
+    (λ _ p → trans (cong (λ eq → resp eq p) ≡⇒≃-refl) (resp-id p)) p
+
+-- A variant of resp-is-equivalence.
+
+resp-is-equivalence′ :
+  ∀ {a p r} {A : Set a}
+  (P : A → Set p) (R : A → A → Set r)
+  (≡↠R : ∀ {x y} → (x ≡ y) ↠ R x y)
+  (resp : ∀ {x y} → R x y → P x → P y) →
+  (∀ x p → resp (_↠_.to ≡↠R (refl x)) p ≡ p) →
+  ∀ {x y} (r : R x y) → Is-equivalence (resp r)
+resp-is-equivalence′ P R ≡↠R resp hyp r =
+  Eq.respects-extensional-equality
+    (λ p → sym $ transport-theorem′ P R ≡↠R resp hyp r p)
+    (subst-is-equivalence P (_↠_.from ≡↠R r))
+
+-- A lemma relating ≃⇒≡, →-cong and cong₂.
+
+≃⇒≡-→-cong :
+  ∀ {ℓ} {A₁ A₂ B₁ B₂ : Set ℓ}
+  (ext : Extensionality ℓ ℓ) →
+  (univ : Univalence ℓ)
+  (A₁≃A₂ : A₁ ≃ A₂) (B₁≃B₂ : B₁ ≃ B₂) →
+  ≃⇒≡ univ (→-cong ext A₁≃A₂ B₁≃B₂) ≡
+    cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂) (≃⇒≡ univ B₁≃B₂)
+≃⇒≡-→-cong {A₂ = A₂} {B₁} ext univ A₁≃A₂ B₁≃B₂ =
+  ≃⇒≡ univ (→-cong ext A₁≃A₂ B₁≃B₂)                        ≡⟨ cong (≃⇒≡ univ) (lift-equality ext lemma) ⟩
+
+  ≃⇒≡ univ (≡⇒≃ (cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂)
+                                       (≃⇒≡ univ B₁≃B₂)))  ≡⟨ left-inverse-of (≡≃≃ univ) _ ⟩∎
+
+  cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂) (≃⇒≡ univ B₁≃B₂)  ∎
+  where
+  open _≃_
+
+  lemma :
+    (λ f → to B₁≃B₂ ∘ f ∘ from A₁≃A₂) ≡
+    to (≡⇒≃ (cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂)
+                                   (≃⇒≡ univ B₁≃B₂)))
+  lemma =
+    (λ f → to B₁≃B₂ ∘ f ∘ from A₁≃A₂)                  ≡⟨ ext (λ _ → transport-theorem (λ B → A₂ → B) (λ A≃B g → _≃_.to A≃B ∘ g)
+                                                                                       refl univ B₁≃B₂ _) ⟩
+    subst (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂) ∘
+    (λ f → f ∘ from A₁≃A₂)                             ≡⟨ cong (_∘_ (subst (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂))) (ext λ f →
+                                                            transport-theorem (λ A → A → B₁) (λ A≃B g → g ∘ _≃_.from A≃B) refl univ A₁≃A₂ f) ⟩
+    subst (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂) ∘
+    subst (λ A → A → B₁) (≃⇒≡ univ A₁≃A₂)              ≡⟨ cong₂ (λ g h f → g (h f)) (ext $ subst-in-terms-of-≡⇒↝ equivalence _ (λ B → A₂ → B))
+                                                                                    (ext $ subst-in-terms-of-≡⇒↝ equivalence _ (λ A → A → B₁)) ⟩
+    to (≡⇒≃ (cong (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂))) ∘
+    to (≡⇒≃ (cong (λ A → A → B₁) (≃⇒≡ univ A₁≃A₂)))    ≡⟨⟩
+
+    to (≡⇒≃ (cong (λ B → A₂ → B) (≃⇒≡ univ B₁≃B₂)) ⊚
+        ≡⇒≃ (cong (λ A → A → B₁) (≃⇒≡ univ A₁≃A₂)))    ≡⟨ cong to $ sym $ ≡⇒≃-trans ext _ _ ⟩∎
+
+    to (≡⇒≃ (cong₂ (λ A B → A → B) (≃⇒≡ univ A₁≃A₂)
+                                   (≃⇒≡ univ B₁≃B₂)))  ∎
+
+-- One can sometimes push cong through ≃⇒≡ (assuming
+-- extensionality).
+
+cong-≃⇒≡ :
+  ∀ {ℓ p} {A B : Set ℓ} {A≃B : A ≃ B} →
+  Extensionality p p →
+  (univ₁ : Univalence ℓ)
+  (univ₂ : Univalence p)
+  (P : Set ℓ → Set p)
+  (P-cong : ∀ {A B} → A ≃ B → P A ≃ P B) →
+  (∀ {A} (p : P A) → _≃_.to (P-cong Eq.id) p ≡ p) →
+  cong P (≃⇒≡ univ₁ A≃B) ≡ ≃⇒≡ univ₂ (P-cong A≃B)
+cong-≃⇒≡ {A≃B = A≃B} ext univ₁ univ₂ P P-cong P-cong-id =
+  cong P (≃⇒≡ univ₁ A≃B)                    ≡⟨ sym $ _≃_.left-inverse-of (≡≃≃ univ₂) _ ⟩
+  ≃⇒≡ univ₂ (≡⇒≃ (cong P (≃⇒≡ univ₁ A≃B)))  ≡⟨ cong (≃⇒≡ univ₂) $ lift-equality ext lemma ⟩∎
+  ≃⇒≡ univ₂ (P-cong A≃B)                    ∎
+  where
+  lemma : ≡⇒→ (cong P (≃⇒≡ univ₁ A≃B)) ≡ _≃_.to (P-cong A≃B)
+  lemma = ext λ x →
+    ≡⇒→ (cong P (≃⇒≡ univ₁ A≃B)) x  ≡⟨ sym $ subst-in-terms-of-≡⇒↝ equivalence _ P x ⟩
+    subst P (≃⇒≡ univ₁ A≃B) x       ≡⟨ sym $ transport-theorem P (_≃_.to ∘ P-cong) P-cong-id univ₁ A≃B x ⟩∎
+    _≃_.to (P-cong A≃B) x           ∎
+
+-- Any "resp" function that preserves identity also preserves
+-- compositions (assuming univalence and extensionality).
+
+resp-preserves-compositions :
+  ∀ {p₁ p₂} (P : Set p₁ → Set p₂) →
+  (resp : ∀ {A B} → A ≃ B → P A → P B) →
+  (∀ {A} (p : P A) → resp Eq.id p ≡ p) →
+  Univalence p₁ → Extensionality p₁ p₁ →
+  ∀ {A B C} (A≃B : A ≃ B) (B≃C : B ≃ C) p →
+  resp (B≃C ⊚ A≃B) p ≡ (resp B≃C ∘ resp A≃B) p
+resp-preserves-compositions P resp resp-id univ ext A≃B B≃C p =
+  resp (B≃C ⊚ A≃B) p                                 ≡⟨ transport-theorem P resp resp-id univ _ _ ⟩
+  subst P (≃⇒≡ univ (B≃C ⊚ A≃B)) p                   ≡⟨ cong (λ eq → subst P eq p) $ ≃⇒≡-∘ univ ext A≃B B≃C ⟩
+  subst P (trans (≃⇒≡ univ A≃B) (≃⇒≡ univ B≃C)) p    ≡⟨ sym $ subst-subst P _ _ _ ⟩
+  subst P (≃⇒≡ univ B≃C) (subst P (≃⇒≡ univ A≃B) p)  ≡⟨ sym $ transport-theorem P resp resp-id univ _ _ ⟩
+  resp B≃C (subst P (≃⇒≡ univ A≃B) p)                ≡⟨ sym $ cong (resp _) $ transport-theorem P resp resp-id univ _ _ ⟩∎
+  resp B≃C (resp A≃B p)                              ∎
+
+-- Any "resp" function that preserves identity also preserves
+-- inverses (assuming univalence and extensionality).
+
+resp-preserves-inverses :
+  ∀ {p₁ p₂} (P : Set p₁ → Set p₂) →
+  (resp : ∀ {A B} → A ≃ B → P A → P B) →
+  (∀ {A} (p : P A) → resp Eq.id p ≡ p) →
+  Univalence p₁ → Extensionality p₁ p₁ →
+  ∀ {A B} (A≃B : A ≃ B) p q →
+  resp A≃B p ≡ q → resp (inverse A≃B) q ≡ p
+resp-preserves-inverses P resp resp-id univ ext A≃B p q eq =
+  let lemma =
+        q                                     ≡⟨ sym eq ⟩
+        resp A≃B p                            ≡⟨ transport-theorem P resp resp-id univ _ _ ⟩
+        subst P (≃⇒≡ univ A≃B) p              ≡⟨ cong (λ eq → subst P eq p) $ sym $ sym-sym _ ⟩∎
+        subst P (sym (sym (≃⇒≡ univ A≃B))) p  ∎
+  in
+
+  resp (inverse A≃B) q                                                 ≡⟨ transport-theorem P resp resp-id univ _ _ ⟩
+  subst P (≃⇒≡ univ (inverse A≃B)) q                                   ≡⟨ cong (λ eq → subst P eq q) $ ≃⇒≡-inverse univ ext A≃B ⟩
+  subst P (sym (≃⇒≡ univ A≃B)) q                                       ≡⟨ cong (subst P (sym (≃⇒≡ univ A≃B))) lemma ⟩
+  subst P (sym (≃⇒≡ univ A≃B)) (subst P (sym (sym (≃⇒≡ univ A≃B))) p)  ≡⟨ subst-subst-sym P _ _ ⟩∎
+  p                                                                    ∎
hunk ./Univalence-axiom.agda 708
-  abstract
-
-    to∘from : ∀ eq → to (from eq) ≡ eq
-    to∘from A₂≡B₂ =
-      let ext₂ = lower-extensionality ℓ₁ ℓ₁ ext in
-
-      _≃_.to (≃-≡ (≡≃≃ univ₂)) (lift-equality ext₂ (
-
-        ≡⇒→ (≃⇒≡ univ₂ ((B₁≃B₂ ⊚ ≡⇒≃ (≃⇒≡ univ₁ ((inverse B₁≃B₂ ⊚
-                             ≡⇒≃ A₂≡B₂) ⊚ A₁≃A₂))) ⊚ inverse A₁≃A₂))  ≡⟨ cong _≃_.to $ _≃_.right-inverse-of (≡≃≃ univ₂) _ ⟩
-
-        (_≃_.to B₁≃B₂ ∘
-         ≡⇒→ (≃⇒≡ univ₁ ((inverse B₁≃B₂ ⊚ ≡⇒≃ A₂≡B₂) ⊚ A₁≃A₂)) ∘
-         _≃_.from A₁≃A₂)                                              ≡⟨ cong (λ eq → _≃_.to B₁≃B₂ ∘ _≃_.to eq ∘ _≃_.from A₁≃A₂) $
-                                                                           _≃_.right-inverse-of (≡≃≃ univ₁) _ ⟩
-        ((_≃_.to B₁≃B₂ ∘ _≃_.from B₁≃B₂) ∘ ≡⇒→ A₂≡B₂ ∘
-         (_≃_.to A₁≃A₂ ∘ _≃_.from A₁≃A₂))                             ≡⟨ cong₂ (λ f g → f ∘ ≡⇒→ A₂≡B₂ ∘ g)
-                                                                               (ext₂ $ _≃_.right-inverse-of B₁≃B₂)
-                                                                               (ext₂ $ _≃_.right-inverse-of A₁≃A₂) ⟩∎
-        ≡⇒→ A₂≡B₂                                                     ∎))
-
-    from∘to : ∀ eq → from (to eq) ≡ eq
-    from∘to A₁≡B₁ =
-      let ext₁ = lower-extensionality ℓ₂ ℓ₂ ext in
-
-      _≃_.to (≃-≡ (≡≃≃ univ₁)) (lift-equality ext₁ (
-
-        ≡⇒→ (≃⇒≡ univ₁ ((inverse B₁≃B₂ ⊚ ≡⇒≃ (≃⇒≡ univ₂ ((B₁≃B₂ ⊚
-                             ≡⇒≃ A₁≡B₁) ⊚ inverse A₁≃A₂))) ⊚ A₁≃A₂))  ≡⟨ cong _≃_.to $ _≃_.right-inverse-of (≡≃≃ univ₁) _ ⟩
-
-        (_≃_.from B₁≃B₂ ∘
-         ≡⇒→ (≃⇒≡ univ₂ ((B₁≃B₂ ⊚ ≡⇒≃ A₁≡B₁) ⊚ inverse A₁≃A₂)) ∘
-         _≃_.to A₁≃A₂)                                                ≡⟨ cong (λ eq → _≃_.from B₁≃B₂ ∘ _≃_.to eq ∘ _≃_.to A₁≃A₂) $
-                                                                           _≃_.right-inverse-of (≡≃≃ univ₂) _ ⟩
-        ((_≃_.from B₁≃B₂ ∘ _≃_.to B₁≃B₂) ∘ ≡⇒→ A₁≡B₁ ∘
-         (_≃_.from A₁≃A₂ ∘ _≃_.to A₁≃A₂))                             ≡⟨ cong₂ (λ f g → f ∘ ≡⇒→ A₁≡B₁ ∘ g)
-                                                                               (ext₁ $ _≃_.left-inverse-of B₁≃B₂)
-                                                                               (ext₁ $ _≃_.left-inverse-of A₁≃A₂) ⟩∎
-        ≡⇒→ A₁≡B₁                                                     ∎))
+  to∘from : ∀ eq → to (from eq) ≡ eq
+  to∘from A₂≡B₂ =
+    let ext₂ = lower-extensionality ℓ₁ ℓ₁ ext in
+
+    _≃_.to (≃-≡ (≡≃≃ univ₂)) (lift-equality ext₂ (
+
+      ≡⇒→ (≃⇒≡ univ₂ ((B₁≃B₂ ⊚ ≡⇒≃ (≃⇒≡ univ₁ ((inverse B₁≃B₂ ⊚
+                           ≡⇒≃ A₂≡B₂) ⊚ A₁≃A₂))) ⊚ inverse A₁≃A₂))  ≡⟨ cong _≃_.to $ _≃_.right-inverse-of (≡≃≃ univ₂) _ ⟩
+
+      (_≃_.to B₁≃B₂ ∘
+       ≡⇒→ (≃⇒≡ univ₁ ((inverse B₁≃B₂ ⊚ ≡⇒≃ A₂≡B₂) ⊚ A₁≃A₂)) ∘
+       _≃_.from A₁≃A₂)                                              ≡⟨ cong (λ eq → _≃_.to B₁≃B₂ ∘ _≃_.to eq ∘ _≃_.from A₁≃A₂) $
+                                                                         _≃_.right-inverse-of (≡≃≃ univ₁) _ ⟩
+      ((_≃_.to B₁≃B₂ ∘ _≃_.from B₁≃B₂) ∘ ≡⇒→ A₂≡B₂ ∘
+       (_≃_.to A₁≃A₂ ∘ _≃_.from A₁≃A₂))                             ≡⟨ cong₂ (λ f g → f ∘ ≡⇒→ A₂≡B₂ ∘ g)
+                                                                             (ext₂ $ _≃_.right-inverse-of B₁≃B₂)
+                                                                             (ext₂ $ _≃_.right-inverse-of A₁≃A₂) ⟩∎
+      ≡⇒→ A₂≡B₂                                                     ∎))
+
+  from∘to : ∀ eq → from (to eq) ≡ eq
+  from∘to A₁≡B₁ =
+    let ext₁ = lower-extensionality ℓ₂ ℓ₂ ext in
+
+    _≃_.to (≃-≡ (≡≃≃ univ₁)) (lift-equality ext₁ (
+
+      ≡⇒→ (≃⇒≡ univ₁ ((inverse B₁≃B₂ ⊚ ≡⇒≃ (≃⇒≡ univ₂ ((B₁≃B₂ ⊚
+                           ≡⇒≃ A₁≡B₁) ⊚ inverse A₁≃A₂))) ⊚ A₁≃A₂))  ≡⟨ cong _≃_.to $ _≃_.right-inverse-of (≡≃≃ univ₁) _ ⟩
+
+      (_≃_.from B₁≃B₂ ∘
+       ≡⇒→ (≃⇒≡ univ₂ ((B₁≃B₂ ⊚ ≡⇒≃ A₁≡B₁) ⊚ inverse A₁≃A₂)) ∘
+       _≃_.to A₁≃A₂)                                                ≡⟨ cong (λ eq → _≃_.from B₁≃B₂ ∘ _≃_.to eq ∘ _≃_.to A₁≃A₂) $
+                                                                         _≃_.right-inverse-of (≡≃≃ univ₂) _ ⟩
+      ((_≃_.from B₁≃B₂ ∘ _≃_.to B₁≃B₂) ∘ ≡⇒→ A₁≡B₁ ∘
+       (_≃_.from A₁≃A₂ ∘ _≃_.to A₁≃A₂))                             ≡⟨ cong₂ (λ f g → f ∘ ≡⇒→ A₁≡B₁ ∘ g)
+                                                                             (ext₁ $ _≃_.left-inverse-of B₁≃B₂)
+                                                                             (ext₁ $ _≃_.left-inverse-of A₁≃A₂) ⟩∎
+      ≡⇒→ A₁≡B₁                                                     ∎))
hunk ./Univalence-axiom/Isomorphism-implies-equality.agda 45
-abstract
-
-  -- If _↔_.to m is a morphism, then _↔_.from m is also a morphism.
-
-  from-also-_-ary-morphism :
-    (n : ℕ) {A B : Set} (f₁ : A ^ n ⟶ A) (f₂ : B ^ n ⟶ B) (m : A ↔ B) →
-    Is- n -ary-morphism f₁ f₂ (_↔_.to m) →
-    Is- n -ary-morphism f₂ f₁ (_↔_.from m)
-  from-also- zero  -ary-morphism f₁ f₂ m is = _↔_.to-from m is
-  from-also- suc n -ary-morphism f₁ f₂ m is = λ x →
-    from-also- n -ary-morphism (f₁ (from x)) (f₂ x) m
-      (subst (λ y → Is- n -ary-morphism (f₁ (from x)) (f₂ y) to)
-             (right-inverse-of x)
-             (is (from x)))
-    where open _↔_ m
+-- If _↔_.to m is a morphism, then _↔_.from m is also a morphism.
+
+from-also-_-ary-morphism :
+  (n : ℕ) {A B : Set} (f₁ : A ^ n ⟶ A) (f₂ : B ^ n ⟶ B) (m : A ↔ B) →
+  Is- n -ary-morphism f₁ f₂ (_↔_.to m) →
+  Is- n -ary-morphism f₂ f₁ (_↔_.from m)
+from-also- zero  -ary-morphism f₁ f₂ m is = _↔_.to-from m is
+from-also- suc n -ary-morphism f₁ f₂ m is = λ x →
+  from-also- n -ary-morphism (f₁ (from x)) (f₂ x) m
+    (subst (λ y → Is- n -ary-morphism (f₁ (from x)) (f₂ y) to)
+           (right-inverse-of x)
+           (is (from x)))
+  where open _↔_ m
hunk ./Univalence-axiom/Isomorphism-implies-equality.agda 65
-abstract
-
-  -- Cast simplification lemma.
-
-  cast-id : {A : Set} →
-            (∀ n → Extensionality′ A (λ _ → A ^ n ⟶ A)) →
-            ∀ n (f : A ^ n ⟶ A) → cast Eq.id n f ≡ f
-  cast-id ext zero    f = refl f
-  cast-id ext (suc n) f = ext n $ λ x → cast-id ext n (f x)
-
-  -- We can express cast as an instance of subst (assuming
-  -- extensionality and univalence).
-
-  cast-is-subst :
-    (∀ {A : Set} n → Extensionality′ A (λ _ → A ^ n ⟶ A)) →
-    {A₁ A₂ : Set}
-    (univ : Univalence′ A₁ A₂)
-    (A₁≃A₂ : A₁ ≃ A₂) (n : ℕ) (f : A₁ ^ n ⟶ A₁) →
-    cast A₁≃A₂ n f ≡ subst (λ C → C ^ n ⟶ C) (≃⇒≡ univ A₁≃A₂) f
-  cast-is-subst ext univ A₁≃A₂ n =
-    transport-theorem
-      (λ A → A ^ n ⟶ A)
-      (λ A≃B f → cast A≃B n f)
-      (cast-id ext n)
-      univ
-      A₁≃A₂
-
-  -- If there is an isomorphism from f₁ to f₂, then the corresponding
-  -- instance of cast maps f₁ to f₂ (assuming extensionality).
-
-  cast-isomorphism :
-    {A₁ A₂ : Set} →
-    (∀ n → Extensionality′ A₂ (λ _ → A₂ ^ n ⟶ A₂)) →
-    (A₁≃A₂ : A₁ ≃ A₂)
-    (n : ℕ) (f₁ : A₁ ^ n ⟶ A₁) (f₂ : A₂ ^ n ⟶ A₂) →
-    Is- n -ary-morphism f₁ f₂ (_≃_.to A₁≃A₂) →
-    cast A₁≃A₂ n f₁ ≡ f₂
-  cast-isomorphism ext A₁≃A₂ zero    f₁ f₂ is = is
-  cast-isomorphism ext A₁≃A₂ (suc n) f₁ f₂ is = ext n $ λ x →
-    cast A₁≃A₂ n (f₁ (from x))  ≡⟨ cast-isomorphism ext A₁≃A₂ n _ _ (is (from x)) ⟩
-    f₂ (to (from x))            ≡⟨ cong f₂ (right-inverse-of x) ⟩∎
-    f₂ x                        ∎
-    where open _≃_ A₁≃A₂
-
-  -- Combining the results above we get the following: if there is an
-  -- isomorphism from f₁ to f₂, then the corresponding instance of
-  -- subst maps f₁ to f₂ (assuming extensionality and univalence).
-
-  subst-isomorphism :
-    (∀ {A : Set} n → Extensionality′ A (λ _ → A ^ n ⟶ A)) →
-    {A₁ A₂ : Set}
-    (univ : Univalence′ A₁ A₂)
-    (A₁≃A₂ : A₁ ≃ A₂)
-    (n : ℕ) (f₁ : A₁ ^ n ⟶ A₁) (f₂ : A₂ ^ n ⟶ A₂) →
-    Is- n -ary-morphism f₁ f₂ (_≃_.to A₁≃A₂) →
-    subst (λ A → A ^ n ⟶ A) (≃⇒≡ univ A₁≃A₂) f₁ ≡ f₂
-  subst-isomorphism ext univ A₁≃A₂ n f₁ f₂ is =
-    subst (λ A → A ^ n ⟶ A) (≃⇒≡ univ A₁≃A₂) f₁  ≡⟨ sym $ cast-is-subst ext univ A₁≃A₂ n f₁ ⟩
-    cast A₁≃A₂ n f₁                              ≡⟨ cast-isomorphism ext A₁≃A₂ n f₁ f₂ is ⟩∎
-    f₂                                           ∎
+-- Cast simplification lemma.
+
+cast-id : {A : Set} →
+          (∀ n → Extensionality′ A (λ _ → A ^ n ⟶ A)) →
+          ∀ n (f : A ^ n ⟶ A) → cast Eq.id n f ≡ f
+cast-id ext zero    f = refl f
+cast-id ext (suc n) f = ext n $ λ x → cast-id ext n (f x)
+
+-- We can express cast as an instance of subst (assuming
+-- extensionality and univalence).
+
+cast-is-subst :
+  (∀ {A : Set} n → Extensionality′ A (λ _ → A ^ n ⟶ A)) →
+  {A₁ A₂ : Set}
+  (univ : Univalence′ A₁ A₂)
+  (A₁≃A₂ : A₁ ≃ A₂) (n : ℕ) (f : A₁ ^ n ⟶ A₁) →
+  cast A₁≃A₂ n f ≡ subst (λ C → C ^ n ⟶ C) (≃⇒≡ univ A₁≃A₂) f
+cast-is-subst ext univ A₁≃A₂ n =
+  transport-theorem
+    (λ A → A ^ n ⟶ A)
+    (λ A≃B f → cast A≃B n f)
+    (cast-id ext n)
+    univ
+    A₁≃A₂
+
+-- If there is an isomorphism from f₁ to f₂, then the corresponding
+-- instance of cast maps f₁ to f₂ (assuming extensionality).
+
+cast-isomorphism :
+  {A₁ A₂ : Set} →
+  (∀ n → Extensionality′ A₂ (λ _ → A₂ ^ n ⟶ A₂)) →
+  (A₁≃A₂ : A₁ ≃ A₂)
+  (n : ℕ) (f₁ : A₁ ^ n ⟶ A₁) (f₂ : A₂ ^ n ⟶ A₂) →
+  Is- n -ary-morphism f₁ f₂ (_≃_.to A₁≃A₂) →
+  cast A₁≃A₂ n f₁ ≡ f₂
+cast-isomorphism ext A₁≃A₂ zero    f₁ f₂ is = is
+cast-isomorphism ext A₁≃A₂ (suc n) f₁ f₂ is = ext n $ λ x →
+  cast A₁≃A₂ n (f₁ (from x))  ≡⟨ cast-isomorphism ext A₁≃A₂ n _ _ (is (from x)) ⟩
+  f₂ (to (from x))            ≡⟨ cong f₂ (right-inverse-of x) ⟩∎
+  f₂ x                        ∎
+  where open _≃_ A₁≃A₂
+
+-- Combining the results above we get the following: if there is an
+-- isomorphism from f₁ to f₂, then the corresponding instance of
+-- subst maps f₁ to f₂ (assuming extensionality and univalence).
+
+subst-isomorphism :
+  (∀ {A : Set} n → Extensionality′ A (λ _ → A ^ n ⟶ A)) →
+  {A₁ A₂ : Set}
+  (univ : Univalence′ A₁ A₂)
+  (A₁≃A₂ : A₁ ≃ A₂)
+  (n : ℕ) (f₁ : A₁ ^ n ⟶ A₁) (f₂ : A₂ ^ n ⟶ A₂) →
+  Is- n -ary-morphism f₁ f₂ (_≃_.to A₁≃A₂) →
+  subst (λ A → A ^ n ⟶ A) (≃⇒≡ univ A₁≃A₂) f₁ ≡ f₂
+subst-isomorphism ext univ A₁≃A₂ n f₁ f₂ is =
+  subst (λ A → A ^ n ⟶ A) (≃⇒≡ univ A₁≃A₂) f₁  ≡⟨ sym $ cast-is-subst ext univ A₁≃A₂ n f₁ ⟩
+  cast A₁≃A₂ n f₁                              ≡⟨ cast-isomorphism ext A₁≃A₂ n f₁ f₂ is ⟩∎
+  f₂                                           ∎
hunk ./Univalence-axiom/Isomorphism-implies-equality.agda 177
-abstract
-
-  -- If _↔_.to m is a morphism, then _↔_.from m is also a morphism.
-
-  from-also-structure-morphism :
-    (s : Structure) →
-    {A B : Set} {s₁ : ⟦ s ⟧ A} {s₂ : ⟦ s ⟧ B} →
-    (m : A ↔ B) →
-    Is-structure-morphism s s₁ s₂ (_↔_.to m) →
-    Is-structure-morphism s s₂ s₁ (_↔_.from m)
-  from-also-structure-morphism empty           m = _
-  from-also-structure-morphism (s +axiom _)    m =
-    from-also-structure-morphism s m
-  from-also-structure-morphism (s +operator n) m =
-    Σ-map (from-also-structure-morphism s m)
-          (from-also- n -ary-morphism _ _ m)
-
-  -- Isomorphic structures are equal (assuming univalence).
-
-  isomorphic-equal :
-    Univalence′ (Set ²/≡) Set →
-    Univalence lzero →
-    (s : Structure) (s₁ s₂ : ⟪ s ⟫) →
-    Isomorphism s s₁ s₂ → s₁ ≡ s₂
-  isomorphic-equal univ₁ univ₂
-    s (A₁ , s₁) (A₂ , s₂) (m , is) =
-
-    (A₁ , s₁)  ≡⟨ Σ-≡,≡→≡ A₁≡A₂ (lemma s s₁ s₂ is) ⟩∎
-    (A₂ , s₂)  ∎
-
-    where
-    open _↔_ m
-
-    -- Extensionality follows from univalence.
-
-    ext : Extensionality lzero lzero
-    ext = dependent-extensionality univ₁ (λ _ → univ₂)
-
-    -- The presence of the bijection implies that the structure's
-    -- underlying types are equal (due to univalence).
-
-    A₁≡A₂ : A₁ ≡ A₂
-    A₁≡A₂ = _≃_.from (≡≃≃ univ₂) $ ↔⇒≃ m
-
-    -- We can lift subst-isomorphism to structures by recursion on
-    -- structure codes.
-
-    lemma : (s : Structure)
-            (s₁ : ⟦ s ⟧ A₁) (s₂ : ⟦ s ⟧ A₂) →
-            Is-structure-morphism s s₁ s₂ to →
-            subst ⟦ s ⟧ A₁≡A₂ s₁ ≡ s₂
-    lemma empty _ _ _ = refl _
-
-    lemma (s +axiom (P , P-prop)) (s₁ , ax₁) (s₂ , ax₂) is =
-      subst (λ A → Σ (⟦ s ⟧ A) (P A)) A₁≡A₂ (s₁ , ax₁)  ≡⟨ push-subst-pair′ ⟦ s ⟧ (uncurry P) (lemma s s₁ s₂ is) ⟩
-      (s₂ , _)                                          ≡⟨ cong (_,_ s₂) $ _⇔_.to propositional⇔irrelevant (P-prop _ _) _ _ ⟩∎
-      (s₂ , ax₂)                                        ∎
-
-    lemma (s +operator n) (s₁ , op₁) (s₂ , op₂) (is-s , is-o) =
-      subst (λ A → ⟦ s ⟧ A × (A ^ n ⟶ A)) A₁≡A₂ (s₁ , op₁)  ≡⟨ push-subst-pair′ ⟦ s ⟧ (λ { (A , _) → A ^ n ⟶ A }) (lemma s s₁ s₂ is-s) ⟩
-
-      (s₂ , subst₂ (λ { (A , _) → A ^ n ⟶ A }) A₁≡A₂
-                   (lemma s s₁ s₂ is-s) op₁)                ≡⟨ cong (_,_ s₂) $ subst₂-proj₁ (λ A → A ^ n ⟶ A) ⟩
-
-      (s₂ , subst (λ A → A ^ n ⟶ A) A₁≡A₂ op₁)              ≡⟨ cong (_,_ s₂) $ subst-isomorphism (λ _ → ext) univ₂ (↔⇒≃ m) n op₁ op₂ is-o ⟩∎
-      (s₂ , op₂)                                            ∎
+-- If _↔_.to m is a morphism, then _↔_.from m is also a morphism.
+
+from-also-structure-morphism :
+  (s : Structure) →
+  {A B : Set} {s₁ : ⟦ s ⟧ A} {s₂ : ⟦ s ⟧ B} →
+  (m : A ↔ B) →
+  Is-structure-morphism s s₁ s₂ (_↔_.to m) →
+  Is-structure-morphism s s₂ s₁ (_↔_.from m)
+from-also-structure-morphism empty           m = _
+from-also-structure-morphism (s +axiom _)    m =
+  from-also-structure-morphism s m
+from-also-structure-morphism (s +operator n) m =
+  Σ-map (from-also-structure-morphism s m)
+        (from-also- n -ary-morphism _ _ m)
+
+-- Isomorphic structures are equal (assuming univalence).
+
+isomorphic-equal :
+  Univalence′ (Set ²/≡) Set →
+  Univalence lzero →
+  (s : Structure) (s₁ s₂ : ⟪ s ⟫) →
+  Isomorphism s s₁ s₂ → s₁ ≡ s₂
+isomorphic-equal univ₁ univ₂
+  s (A₁ , s₁) (A₂ , s₂) (m , is) =
+
+  (A₁ , s₁)  ≡⟨ Σ-≡,≡→≡ A₁≡A₂ (lemma s s₁ s₂ is) ⟩∎
+  (A₂ , s₂)  ∎
+
+  where
+  open _↔_ m
+
+  -- Extensionality follows from univalence.
+
+  ext : Extensionality lzero lzero
+  ext = dependent-extensionality univ₁ (λ _ → univ₂)
+
+  -- The presence of the bijection implies that the structure's
+  -- underlying types are equal (due to univalence).
+
+  A₁≡A₂ : A₁ ≡ A₂
+  A₁≡A₂ = _≃_.from (≡≃≃ univ₂) $ ↔⇒≃ m
+
+  -- We can lift subst-isomorphism to structures by recursion on
+  -- structure codes.
+
+  lemma : (s : Structure)
+          (s₁ : ⟦ s ⟧ A₁) (s₂ : ⟦ s ⟧ A₂) →
+          Is-structure-morphism s s₁ s₂ to →
+          subst ⟦ s ⟧ A₁≡A₂ s₁ ≡ s₂
+  lemma empty _ _ _ = refl _
+
+  lemma (s +axiom (P , P-prop)) (s₁ , ax₁) (s₂ , ax₂) is =
+    subst (λ A → Σ (⟦ s ⟧ A) (P A)) A₁≡A₂ (s₁ , ax₁)  ≡⟨ push-subst-pair′ ⟦ s ⟧ (uncurry P) (lemma s s₁ s₂ is) ⟩
+    (s₂ , _)                                          ≡⟨ cong (_,_ s₂) $ _⇔_.to propositional⇔irrelevant (P-prop _ _) _ _ ⟩∎
+    (s₂ , ax₂)                                        ∎
+
+  lemma (s +operator n) (s₁ , op₁) (s₂ , op₂) (is-s , is-o) =
+    subst (λ A → ⟦ s ⟧ A × (A ^ n ⟶ A)) A₁≡A₂ (s₁ , op₁)  ≡⟨ push-subst-pair′ ⟦ s ⟧ (λ { (A , _) → A ^ n ⟶ A }) (lemma s s₁ s₂ is-s) ⟩
+
+    (s₂ , subst₂ (λ { (A , _) → A ^ n ⟶ A }) A₁≡A₂
+                 (lemma s s₁ s₂ is-s) op₁)                ≡⟨ cong (_,_ s₂) $ subst₂-proj₁ (λ A → A ^ n ⟶ A) ⟩
+
+    (s₂ , subst (λ A → A ^ n ⟶ A) A₁≡A₂ op₁)              ≡⟨ cong (_,_ s₂) $ subst-isomorphism (λ _ → ext) univ₂ (↔⇒≃ m) n op₁ op₂ is-o ⟩∎
+    (s₂ , op₂)                                            ∎
hunk ./Univalence-axiom/Isomorphism-is-equality/More.agda 43
-  abstract
-
-    ext : Extensionality (# 0) (# 0)
-    ext = lower-extensionality (# 1) (# 1) ext₁
+  ext : Extensionality (# 0) (# 0)
+  ext = lower-extensionality (# 1) (# 1) ext₁
hunk ./Univalence-axiom/Isomorphism-is-equality/More.agda 236
-  abstract
-
-    -- Every element is isomorphic to itself, transported along the
-    -- "outer" isomorphism.
-
-    isomorphic-to-itself″ :
-      (ass : Assumptions) →
-      ∀ {I J} (I≅J : Isomorphic ass c I J) {x} →
-      Iso″ ass I≅J x
-           (subst Ext (_≃_.to (isomorphism≃equality ass c) I≅J) x)
-    isomorphic-to-itself″ ass I≅J {x} = transport-theorem′
-      Ext
-      (Isomorphic ass c)
-      (_≃_.surjection $ inverse $ isomorphism≃equality ass c)
-      (resp ass)
-      (λ _ → resp-refl ass)
-      I≅J
-      x
+  -- Every element is isomorphic to itself, transported along the
+  -- "outer" isomorphism.
+
+  isomorphic-to-itself″ :
+    (ass : Assumptions) →
+    ∀ {I J} (I≅J : Isomorphic ass c I J) {x} →
+    Iso″ ass I≅J x
+         (subst Ext (_≃_.to (isomorphism≃equality ass c) I≅J) x)
+  isomorphic-to-itself″ ass I≅J {x} = transport-theorem′
+    Ext
+    (Isomorphic ass c)
+    (_≃_.surjection $ inverse $ isomorphism≃equality ass c)
+    (resp ass)
+    (λ _ → resp-refl ass)
+    I≅J
+    x
hunk ./Univalence-axiom/Isomorphism-is-equality/More.agda 276
-    abstract
-      right-inverse-of : ∀ x → to (from x) ≡ x
-      right-inverse-of = _≃_.right-inverse-of I≃I′
-
-      irrelevance : ∀ y (p : to ⁻¹ y) → (from y , right-inverse-of y) ≡ p
-      irrelevance = _≃_.irrelevance I≃I′
+    right-inverse-of : ∀ x → to (from x) ≡ x
+    right-inverse-of = _≃_.right-inverse-of I≃I′
+
+    irrelevance : ∀ y (p : to ⁻¹ y) → (from y , right-inverse-of y) ≡ p
+    irrelevance = _≃_.irrelevance I≃I′
hunk ./Univalence-axiom/Isomorphism-is-equality/More.agda 297
-  abstract
-
-    -- Simplification lemmas.
-
-    resp-refl-lemma :
-      (ass : Assumptions) →
-      ∀ I x →
-      resp-refl ass x ≡
-      _≃_.from (≡⇒≃ $ cong (λ z → z ≡ x) $
-                      isomorphic-to-itself″ ass (reflexivity ass c I))
-               (subst (λ eq → subst Ext eq x ≡ x)
-                      (sym $ _≃_.right-inverse-of
-                               (isomorphism≃equality ass c)
-                               (refl I))
-                      (refl x))
-    resp-refl-lemma ass I x =
-      let rfl    = reflexivity ass c I
-          iso≃eq = λ {I J} → isomorphism≃equality ass c {I = I} {J = J}
-          rio    = right-inverse-of iso≃eq (refl I)
-          lio    = left-inverse-of (inverse iso≃eq) (refl I)
-          sx≡x   = subst (λ eq → subst Ext eq x ≡ x) (sym rio) (refl x)
-
-          sx≡x-lemma =
-            cong (λ eq → subst Ext eq x) rio                      ≡⟨ sym $ trans-reflʳ _ ⟩
-
-            trans (cong (λ eq → subst Ext eq x) rio)
-                  (refl x)                                        ≡⟨ sym $ subst-trans (cong (λ eq → subst Ext eq x) rio) ⟩
-
-            subst (λ z → z ≡ x)
-                  (sym $ cong (λ eq → subst Ext eq x) rio)
-                  (refl x)                                        ≡⟨ cong (λ eq → subst (λ z → z ≡ x) eq (refl x)) $ sym $
-                                                                          cong-sym (λ eq → subst Ext eq x) rio ⟩
-            subst (λ z → z ≡ x)
-                  (cong (λ eq → subst Ext eq x) $ sym rio)
-                  (refl x)                                        ≡⟨ sym $ subst-∘ (λ z → z ≡ x) (λ eq → subst Ext eq x) (sym rio) ⟩∎
-
-            subst (λ eq → subst Ext eq x ≡ x) (sym rio) (refl x)  ∎
-
-          lemma₁ =
-            trans (sym lio) rio  ≡⟨ cong (λ eq → trans (sym eq) rio) $ left-inverse-of∘inverse iso≃eq ⟩
-
-            trans (sym rio) rio  ≡⟨ trans-symˡ rio ⟩∎
-
-            refl (refl I)        ∎
-
-          lemma₂ =
-            elim-refl (λ {I J} _ → Ext I → Ext J) (λ _ e → e)  ≡⟨⟩
-
-            cong (subst Ext) (refl (refl I))                   ≡⟨ cong (cong (subst Ext)) $ sym lemma₁ ⟩∎
-
-            cong (subst Ext) (trans (sym lio) rio)             ∎
-
-          lemma₃ =
-            cong (λ r → r x)
-                 (elim-refl (λ {I J} _ → Ext I → Ext J) (λ _ e → e))   ≡⟨ cong (cong (λ r → r x)) lemma₂ ⟩
-
-            cong (λ r → r x) (cong (subst Ext) (trans (sym lio) rio))  ≡⟨ cong-∘ (λ r → r x) (subst Ext) (trans (sym lio) rio) ⟩
-
-            cong (λ eq → subst Ext eq x) (trans (sym lio) rio)         ≡⟨ cong-trans (λ eq → subst Ext eq x) (sym lio) rio ⟩
-
-            trans (cong (λ eq → subst Ext eq x) (sym lio))
-                  (cong (λ eq → subst Ext eq x) rio)                   ≡⟨ cong (λ eq → trans eq (cong (λ eq → subst Ext eq x) rio)) $
-                                                                               cong-sym (λ eq → subst Ext eq x) lio ⟩∎
-            trans (sym (cong (λ eq → subst Ext eq x) lio))
-                  (cong (λ eq → subst Ext eq x) rio)                   ∎
-      in
-
-      resp-refl ass x                                                  ≡⟨ sym $ trans-reflʳ _ ⟩
-
-      trans (resp-refl ass x) (refl x)                                 ≡⟨ cong (trans (resp-refl ass x)) $ trans-symˡ (subst-refl Ext x) ⟩
-
-      trans (resp-refl ass x)
-            (trans (sym $ subst-refl Ext x) (subst-refl Ext x))        ≡⟨ sym $ trans-assoc _ (sym $ subst-refl Ext x) (subst-refl Ext x) ⟩
-
-      trans (trans (resp-refl ass x) (sym $ subst-refl Ext x))
-            (subst-refl Ext x)                                         ≡⟨ cong (trans (trans (resp-refl ass x) (sym $ subst-refl Ext x)))
-                                                                               lemma₃ ⟩
-      trans (trans (resp-refl ass x) (sym $ subst-refl Ext x))
-            (trans (sym (cong (λ eq → subst Ext eq x) lio))
-                   (cong (λ eq → subst Ext eq x) rio))                 ≡⟨ sym $ trans-assoc _ _ (cong (λ eq → subst Ext eq x) rio) ⟩
-
-      trans (trans (trans (resp-refl ass x) (sym $ subst-refl Ext x))
-                   (sym (cong (λ eq → subst Ext eq x) lio)))
-            (cong (λ eq → subst Ext eq x) rio)                         ≡⟨ cong₂ trans
-                                                                            (sym $ transport-theorem′-refl Ext (Isomorphic ass c) (inverse iso≃eq)
-                                                                                                           (resp ass) (λ _ → resp-refl ass) x)
-                                                                            sx≡x-lemma ⟩
-      trans (isomorphic-to-itself″ ass rfl) sx≡x                       ≡⟨ sym $ subst-trans (isomorphic-to-itself″ ass rfl) ⟩
-
-      subst (λ z → z ≡ x) (sym $ isomorphic-to-itself″ ass rfl) sx≡x   ≡⟨ subst-in-terms-of-inverse∘≡⇒↝ equivalence
-                                                                            (isomorphic-to-itself″ ass rfl) (λ z → z ≡ x) _ ⟩∎
-      from (≡⇒≃ $ cong (λ z → z ≡ x) $ isomorphic-to-itself″ ass rfl)
-           sx≡x                                                        ∎
-
-      where open _≃_
-
-    isomorphic-to-itself-reflexivity :
-      (ass : Assumptions) →
-      ∀ I x →
-      isomorphic-to-itself ass (reflexivity ass c I) x ≡
-      subst (Iso ass (reflexivity ass c I) x)
-            (sym $ resp-refl ass x)
-            (reflexivityE ass c extension I x)
-    isomorphic-to-itself-reflexivity ass I x =
-      let rfl = reflexivity ass c I
-          r-r = resp-refl ass x in
-
-      from (Iso≃Iso″ ass rfl) (refl (resp ass rfl x))                  ≡⟨ elim¹ (λ {y} resp-x≡y → from (Iso≃Iso″ ass rfl) (refl (resp ass rfl x)) ≡
-                                                                                                  subst (Iso ass rfl x) (sym resp-x≡y)
-                                                                                                        (from (Iso≃Iso″ ass rfl) resp-x≡y))
-                                                                                (refl _) r-r ⟩
-      subst (Iso ass rfl x) (sym r-r) (from (Iso≃Iso″ ass rfl) r-r)    ≡⟨ cong (subst (Iso ass rfl x) (sym r-r) ∘ from (Iso≃Iso″ ass rfl))
-                                                                               (resp-refl-lemma ass I x) ⟩∎
-      subst (Iso ass rfl x) (sym r-r)
-        (from (Iso≃Iso″ ass rfl)
-           (from
-              (≡⇒≃ $ cong (λ z → z ≡ x)
-                          (isomorphic-to-itself″ ass rfl))
-              (subst (λ eq → subst Ext eq x ≡ x)
-                 (sym $ right-inverse-of (isomorphism≃equality ass c)
-                                         (refl I)) (refl x))))         ∎
-
-      where open _≃_
+  -- Simplification lemmas.
+
+  resp-refl-lemma :
+    (ass : Assumptions) →
+    ∀ I x →
+    resp-refl ass x ≡
+    _≃_.from (≡⇒≃ $ cong (λ z → z ≡ x) $
+                    isomorphic-to-itself″ ass (reflexivity ass c I))
+             (subst (λ eq → subst Ext eq x ≡ x)
+                    (sym $ _≃_.right-inverse-of
+                             (isomorphism≃equality ass c)
+                             (refl I))
+                    (refl x))
+  resp-refl-lemma ass I x =
+    let rfl    = reflexivity ass c I
+        iso≃eq = λ {I J} → isomorphism≃equality ass c {I = I} {J = J}
+        rio    = right-inverse-of iso≃eq (refl I)
+        lio    = left-inverse-of (inverse iso≃eq) (refl I)
+        sx≡x   = subst (λ eq → subst Ext eq x ≡ x) (sym rio) (refl x)
+
+        sx≡x-lemma =
+          cong (λ eq → subst Ext eq x) rio                      ≡⟨ sym $ trans-reflʳ _ ⟩
+
+          trans (cong (λ eq → subst Ext eq x) rio)
+                (refl x)                                        ≡⟨ sym $ subst-trans (cong (λ eq → subst Ext eq x) rio) ⟩
+
+          subst (λ z → z ≡ x)
+                (sym $ cong (λ eq → subst Ext eq x) rio)
+                (refl x)                                        ≡⟨ cong (λ eq → subst (λ z → z ≡ x) eq (refl x)) $ sym $
+                                                                        cong-sym (λ eq → subst Ext eq x) rio ⟩
+          subst (λ z → z ≡ x)
+                (cong (λ eq → subst Ext eq x) $ sym rio)
+                (refl x)                                        ≡⟨ sym $ subst-∘ (λ z → z ≡ x) (λ eq → subst Ext eq x) (sym rio) ⟩∎
+
+          subst (λ eq → subst Ext eq x ≡ x) (sym rio) (refl x)  ∎
+
+        lemma₁ =
+          trans (sym lio) rio  ≡⟨ cong (λ eq → trans (sym eq) rio) $ left-inverse-of∘inverse iso≃eq ⟩
+
+          trans (sym rio) rio  ≡⟨ trans-symˡ rio ⟩∎
+
+          refl (refl I)        ∎
+
+        lemma₂ =
+          elim-refl (λ {I J} _ → Ext I → Ext J) (λ _ e → e)  ≡⟨⟩
+
+          cong (subst Ext) (refl (refl I))                   ≡⟨ cong (cong (subst Ext)) $ sym lemma₁ ⟩∎
+
+          cong (subst Ext) (trans (sym lio) rio)             ∎
+
+        lemma₃ =
+          cong (λ r → r x)
+               (elim-refl (λ {I J} _ → Ext I → Ext J) (λ _ e → e))   ≡⟨ cong (cong (λ r → r x)) lemma₂ ⟩
+
+          cong (λ r → r x) (cong (subst Ext) (trans (sym lio) rio))  ≡⟨ cong-∘ (λ r → r x) (subst Ext) (trans (sym lio) rio) ⟩
+
+          cong (λ eq → subst Ext eq x) (trans (sym lio) rio)         ≡⟨ cong-trans (λ eq → subst Ext eq x) (sym lio) rio ⟩
+
+          trans (cong (λ eq → subst Ext eq x) (sym lio))
+                (cong (λ eq → subst Ext eq x) rio)                   ≡⟨ cong (λ eq → trans eq (cong (λ eq → subst Ext eq x) rio)) $
+                                                                             cong-sym (λ eq → subst Ext eq x) lio ⟩∎
+          trans (sym (cong (λ eq → subst Ext eq x) lio))
+                (cong (λ eq → subst Ext eq x) rio)                   ∎
+    in
+
+    resp-refl ass x                                                  ≡⟨ sym $ trans-reflʳ _ ⟩
+
+    trans (resp-refl ass x) (refl x)                                 ≡⟨ cong (trans (resp-refl ass x)) $ trans-symˡ (subst-refl Ext x) ⟩
+
+    trans (resp-refl ass x)
+          (trans (sym $ subst-refl Ext x) (subst-refl Ext x))        ≡⟨ sym $ trans-assoc _ (sym $ subst-refl Ext x) (subst-refl Ext x) ⟩
+
+    trans (trans (resp-refl ass x) (sym $ subst-refl Ext x))
+          (subst-refl Ext x)                                         ≡⟨ cong (trans (trans (resp-refl ass x) (sym $ subst-refl Ext x)))
+                                                                             lemma₃ ⟩
+    trans (trans (resp-refl ass x) (sym $ subst-refl Ext x))
+          (trans (sym (cong (λ eq → subst Ext eq x) lio))
+                 (cong (λ eq → subst Ext eq x) rio))                 ≡⟨ sym $ trans-assoc _ _ (cong (λ eq → subst Ext eq x) rio) ⟩
+
+    trans (trans (trans (resp-refl ass x) (sym $ subst-refl Ext x))
+                 (sym (cong (λ eq → subst Ext eq x) lio)))
+          (cong (λ eq → subst Ext eq x) rio)                         ≡⟨ cong₂ trans
+                                                                          (sym $ transport-theorem′-refl Ext (Isomorphic ass c) (inverse iso≃eq)
+                                                                                                         (resp ass) (λ _ → resp-refl ass) x)
+                                                                          sx≡x-lemma ⟩
+    trans (isomorphic-to-itself″ ass rfl) sx≡x                       ≡⟨ sym $ subst-trans (isomorphic-to-itself″ ass rfl) ⟩
+
+    subst (λ z → z ≡ x) (sym $ isomorphic-to-itself″ ass rfl) sx≡x   ≡⟨ subst-in-terms-of-inverse∘≡⇒↝ equivalence
+                                                                          (isomorphic-to-itself″ ass rfl) (λ z → z ≡ x) _ ⟩∎
+    from (≡⇒≃ $ cong (λ z → z ≡ x) $ isomorphic-to-itself″ ass rfl)
+         sx≡x                                                        ∎
+
+    where open _≃_
+
+  isomorphic-to-itself-reflexivity :
+    (ass : Assumptions) →
+    ∀ I x →
+    isomorphic-to-itself ass (reflexivity ass c I) x ≡
+    subst (Iso ass (reflexivity ass c I) x)
+          (sym $ resp-refl ass x)
+          (reflexivityE ass c extension I x)
+  isomorphic-to-itself-reflexivity ass I x =
+    let rfl = reflexivity ass c I
+        r-r = resp-refl ass x in
+
+    from (Iso≃Iso″ ass rfl) (refl (resp ass rfl x))                  ≡⟨ elim¹ (λ {y} resp-x≡y → from (Iso≃Iso″ ass rfl) (refl (resp ass rfl x)) ≡
+                                                                                                subst (Iso ass rfl x) (sym resp-x≡y)
+                                                                                                      (from (Iso≃Iso″ ass rfl) resp-x≡y))
+                                                                              (refl _) r-r ⟩
+    subst (Iso ass rfl x) (sym r-r) (from (Iso≃Iso″ ass rfl) r-r)    ≡⟨ cong (subst (Iso ass rfl x) (sym r-r) ∘ from (Iso≃Iso″ ass rfl))
+                                                                             (resp-refl-lemma ass I x) ⟩∎
+    subst (Iso ass rfl x) (sym r-r)
+      (from (Iso≃Iso″ ass rfl)
+         (from
+            (≡⇒≃ $ cong (λ z → z ≡ x)
+                        (isomorphic-to-itself″ ass rfl))
+            (subst (λ eq → subst Ext eq x ≡ x)
+               (sym $ right-inverse-of (isomorphism≃equality ass c)
+                                       (refl I)) (refl x))))         ∎
+
+    where open _≃_
hunk ./Univalence-axiom/Isomorphism-is-equality/More.agda 802
-  abstract
-
-    -- Reflexivity is mapped to identity.
-
-    cast-refl ass     set          = refl Eq.id
-    cast-refl ass {c} (base A) {I} =
-
-      Type-cong A ass (reflexivity ass c I)  ≡⟨ Type-cong-reflexivity A ass I ⟩∎
-
-      Eq.id                                  ∎
-
-    cast-refl ass {c} (Π σ τ) {I} =
-      let rfl   = reflexivity ass c I
-          rflE  = reflexivityE ass c (Dep σ) I in
-
-      lift-equality-inverse ext₁ $ ext₁ λ f → ext₁ λ x →
-
-        from (cast ass τ (rfl , isomorphic-to-itself σ ass rfl x))
-             (f (resp σ ass rfl x))                                 ≡⟨ cong (λ iso → from (cast ass τ (rfl , iso)) (f (resp σ ass rfl x))) $
-                                                                            isomorphic-to-itself-reflexivity σ ass I x ⟩
-        from (cast ass τ (rfl , subst (Iso ass σ rfl x)
-                                      (sym $ resp-refl σ ass x)
-                                      (rflE x)))
-             (f (resp σ ass rfl x))                                 ≡⟨ elim¹ (λ {y} x≡y →
-                                                                                from (cast ass τ (rfl , subst (Iso ass σ rfl x)
-                                                                                                              x≡y (rflE x)))
-                                                                                     (f y) ≡
-                                                                                f x)
-                                                                             (cong (λ h → _≃_.from h (f x)) $ cast-refl ass τ)
-                                                                             (sym $ resp-refl σ ass x) ⟩∎
-        f x                                                         ∎
-
-      where
-      open _≃_
-      open Assumptions ass
+  -- Reflexivity is mapped to identity.
+
+  cast-refl ass     set          = refl Eq.id
+  cast-refl ass {c} (base A) {I} =
+
+    Type-cong A ass (reflexivity ass c I)  ≡⟨ Type-cong-reflexivity A ass I ⟩∎
+
+    Eq.id                                  ∎
+
+  cast-refl ass {c} (Π σ τ) {I} =
+    let rfl   = reflexivity ass c I
+        rflE  = reflexivityE ass c (Dep σ) I in
+
+    lift-equality-inverse ext₁ $ ext₁ λ f → ext₁ λ x →
+
+      from (cast ass τ (rfl , isomorphic-to-itself σ ass rfl x))
+           (f (resp σ ass rfl x))                                 ≡⟨ cong (λ iso → from (cast ass τ (rfl , iso)) (f (resp σ ass rfl x))) $
+                                                                          isomorphic-to-itself-reflexivity σ ass I x ⟩
+      from (cast ass τ (rfl , subst (Iso ass σ rfl x)
+                                    (sym $ resp-refl σ ass x)
+                                    (rflE x)))
+           (f (resp σ ass rfl x))                                 ≡⟨ elim¹ (λ {y} x≡y →
+                                                                              from (cast ass τ (rfl , subst (Iso ass σ rfl x)
+                                                                                                            x≡y (rflE x)))
+                                                                                   (f y) ≡
+                                                                              f x)
+                                                                           (cong (λ h → _≃_.from h (f x)) $ cast-refl ass τ)
+                                                                           (sym $ resp-refl σ ass x) ⟩∎
+      f x                                                         ∎
+
+    where
+    open _≃_
+    open Assumptions ass
hunk ./Univalence-axiom/Isomorphism-is-equality/More.agda 884
-  abstract
-
-    -- Two definitions of isomorphism are equivalent.
-
-    Iso≃Iso″ ass σ I≅J {f} {g} =
-      Iso ass σ I≅J f g                  ↝⟨ Iso≃Iso‴ ass σ I≅J ⟩
-
-      (f ≡ _≃_.from (cast ass σ I≅J) g)  ↝⟨ inverse $ from≡↔≡to (inverse $ cast ass σ I≅J) ⟩□
-
-      (_≃_.to (cast ass σ I≅J) f ≡ g)    □
+  -- Two definitions of isomorphism are equivalent.
+
+  Iso≃Iso″ ass σ I≅J {f} {g} =
+    Iso ass σ I≅J f g                  ↝⟨ Iso≃Iso‴ ass σ I≅J ⟩
+
+    (f ≡ _≃_.from (cast ass σ I≅J) g)  ↝⟨ inverse $ from≡↔≡to (inverse $ cast ass σ I≅J) ⟩□
+
+    (_≃_.to (cast ass σ I≅J) f ≡ g)    □
hunk ./Univalence-axiom/Isomorphism-is-equality/More.agda 897
-  abstract
-
-    reflexivityE-set :
-      (ass : Assumptions) →
-      ∀ {c} {I : ⟦ c ⟧} {A} →
-      reflexivityE ass c (Dep set) I A ≡ lift Eq.id
-    reflexivityE-set ass {c} {I} {A} =
-
-      reflexivityE ass c (Dep set) I A                                 ≡⟨⟩
-
-      lift (≡⇒≃ (to (from≡↔≡to (inverse Eq.id))
-              (from (≡⇒≃ $ cong (λ B → B ≡ A) $
-                       isomorphic-to-itself″ set ass
-                         (reflexivity ass c I))
-                    (subst (λ eq → subst (λ _ → Set) eq A ≡ A)
-                           (sym $ right-inverse-of
-                                    (isomorphism≃equality ass c)
-                                    (refl I))
-                           (refl A)))))                                ≡⟨ cong (λ eq → lift (≡⇒≃ (to (from≡↔≡to (inverse Eq.id)) eq))) $ sym $
-                                                                            resp-refl-lemma set ass I A ⟩
-      lift (≡⇒≃ (to (from≡↔≡to (inverse Eq.id))
-                    (resp-refl set ass {I = I} A)))                    ≡⟨⟩
-
-      lift (≡⇒≃ (to (from≡↔≡to (inverse Eq.id)) (refl A)))             ≡⟨⟩
-
-      lift (≡⇒≃ (≡⇒→ (cong (λ B → B ≡ A)
-                           (right-inverse-of (inverse Eq.id) A))
-                     (cong id (refl A))))                              ≡⟨⟩
-
-      lift (≡⇒≃ (≡⇒→ (cong (λ B → B ≡ A) (left-inverse-of Eq.id A))
-                     (cong id (refl A))))                              ≡⟨ cong (λ eq → lift (≡⇒≃ (≡⇒→ (cong (λ B → B ≡ A) eq) (refl A))))
-                                                                          left-inverse-of-id  ⟩
-      lift (≡⇒≃ (≡⇒→ (cong (λ B → B ≡ A) (refl A)) (refl A)))          ≡⟨⟩
-
-      lift (≡⇒≃ (≡⇒→ (refl (A ≡ A)) (refl A)))                         ≡⟨⟩
-
-      lift (≡⇒≃ (refl A))                                              ≡⟨ refl _ ⟩∎
-
-      lift Eq.id                                                       ∎
-
-      where open _≃_
+  reflexivityE-set :
+    (ass : Assumptions) →
+    ∀ {c} {I : ⟦ c ⟧} {A} →
+    reflexivityE ass c (Dep set) I A ≡ lift Eq.id
+  reflexivityE-set ass {c} {I} {A} =
+
+    reflexivityE ass c (Dep set) I A                                 ≡⟨⟩
+
+    lift (≡⇒≃ (to (from≡↔≡to (inverse Eq.id))
+            (from (≡⇒≃ $ cong (λ B → B ≡ A) $
+                     isomorphic-to-itself″ set ass
+                       (reflexivity ass c I))
+                  (subst (λ eq → subst (λ _ → Set) eq A ≡ A)
+                         (sym $ right-inverse-of
+                                  (isomorphism≃equality ass c)
+                                  (refl I))
+                         (refl A)))))                                ≡⟨ cong (λ eq → lift (≡⇒≃ (to (from≡↔≡to (inverse Eq.id)) eq))) $ sym $
+                                                                          resp-refl-lemma set ass I A ⟩
+    lift (≡⇒≃ (to (from≡↔≡to (inverse Eq.id))
+                  (resp-refl set ass {I = I} A)))                    ≡⟨⟩
+
+    lift (≡⇒≃ (to (from≡↔≡to (inverse Eq.id)) (refl A)))             ≡⟨⟩
+
+    lift (≡⇒≃ (≡⇒→ (cong (λ B → B ≡ A)
+                         (right-inverse-of (inverse Eq.id) A))
+                   (cong id (refl A))))                              ≡⟨⟩
+
+    lift (≡⇒≃ (≡⇒→ (cong (λ B → B ≡ A) (left-inverse-of Eq.id A))
+                   (cong id (refl A))))                              ≡⟨ cong (λ eq → lift (≡⇒≃ (≡⇒→ (cong (λ B → B ≡ A) eq) (refl A))))
+                                                                        left-inverse-of-id  ⟩
+    lift (≡⇒≃ (≡⇒→ (cong (λ B → B ≡ A) (refl A)) (refl A)))          ≡⟨⟩
+
+    lift (≡⇒≃ (≡⇒→ (refl (A ≡ A)) (refl A)))                         ≡⟨⟩
+
+    lift (≡⇒≃ (refl A))                                              ≡⟨ refl _ ⟩∎
+
+    lift Eq.id                                                       ∎
+
+    where open _≃_
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple.agda 51
-  abstract
-
-    -- Extensionality.
-
-    ext : ∀ {ℓ} → Extensionality ℓ (# 1)
-    ext = dependent-extensionality univ₂ (λ _ → univ₁)
-
-    ext₁ : Extensionality (# 1) (# 1)
-    ext₁ = ext
+  -- Extensionality.
+
+  ext : ∀ {ℓ} → Extensionality ℓ (# 1)
+  ext = dependent-extensionality univ₂ (λ _ → univ₁)
+
+  ext₁ : Extensionality (# 1) (# 1)
+  ext₁ = ext
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple.agda 165
-  abstract
-
-    -- One can prove that two instances of a structure are equal by
-    -- proving that the carrier types and "elements" (suitably
-    -- transported) are equal (assuming univalence).
-
-    equality-pair-lemma :
-      Assumptions →
-      ∀ c {X Y : Instance c} →
-      (X ≡ Y) ↔
-      ∃ λ (eq : Carrier c X ≡ Carrier c Y) →
-        subst (El (proj₁ c)) eq (element c X) ≡ element c Y
-    equality-pair-lemma ass (a , P) {C , x , p} {D , y , q} =
-
-      ((C , x , p) ≡ (D , y , q))                 ↔⟨ inverse $ ≃-≡ $ ↔⇒≃ Σ-assoc ⟩
-
-      (((C , x) , p) ≡ ((D , y) , q))             ↝⟨ inverse $ ignore-propositional-component (proj₂ (P D y) ass) ⟩
-
-      ((C , x) ≡ (D , y))                         ↝⟨ inverse Σ-≡,≡↔≡ ⟩□
-
-      (∃ λ (eq : C ≡ D) → subst (El a) eq x ≡ y)  □
+  -- One can prove that two instances of a structure are equal by
+  -- proving that the carrier types and "elements" (suitably
+  -- transported) are equal (assuming univalence).
+
+  equality-pair-lemma :
+    Assumptions →
+    ∀ c {X Y : Instance c} →
+    (X ≡ Y) ↔
+    ∃ λ (eq : Carrier c X ≡ Carrier c Y) →
+      subst (El (proj₁ c)) eq (element c X) ≡ element c Y
+  equality-pair-lemma ass (a , P) {C , x , p} {D , y , q} =
+
+    ((C , x , p) ≡ (D , y , q))                 ↔⟨ inverse $ ≃-≡ $ ↔⇒≃ Σ-assoc ⟩
+
+    (((C , x) , p) ≡ ((D , y) , q))             ↝⟨ inverse $ ignore-propositional-component (proj₂ (P D y) ass) ⟩
+
+    ((C , x) ≡ (D , y))                         ↝⟨ inverse Σ-≡,≡↔≡ ⟩□
+
+    (∃ λ (eq : C ≡ D) → subst (El a) eq x ≡ y)  □
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple.agda 224
-  abstract
-
-    -- The type of (lifted) isomorphisms between two instances of a
-    -- structure is equal to the type of equalities between the same
-    -- instances (assuming univalence).
-    --
-    -- In short, isomorphism is equal to equality.
-
-    isomorphic≡≡ :
-      Assumptions →
-      ∀ c {X Y} → ↑ (# 2) (Isomorphic c X Y) ≡ (X ≡ Y)
-    isomorphic≡≡ ass c {X} {Y} =
-      ≃⇒≡ univ₂ $ ↔⇒≃ (
-        ↑ _ (Isomorphic c X Y)  ↝⟨ ↑↔ ⟩
-        Isomorphic c X Y        ↝⟨ isomorphism-is-equality ass c X Y ⟩□
-        (X ≡ Y)                 □)
-      where open Assumptions ass
-
-    -- The "first part" of the from component of
-    -- isomorphism-is-equality is equal to a simple function.
-
-    proj₁-from-isomorphism-is-equality :
-      ∀ ass c X Y →
-      proj₁ ∘ _↔_.from (isomorphism-is-equality ass c X Y) ≡
-      elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id)
-    proj₁-from-isomorphism-is-equality ass _ _ _ = ext λ eq →
-
-      ≡⇒≃ (proj₁ (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡
-        (cong (λ { (x , (y , z)) → (x , y) , z }) eq)))))            ≡⟨ cong (≡⇒≃ ∘ proj₁ ∘ Σ-≡,≡←≡) $ proj₁-Σ-≡,≡←≡ _ ⟩
-
-      ≡⇒≃ (proj₁ (Σ-≡,≡←≡ (cong proj₁
-        (cong (λ { (x , (y , z)) → (x , y) , z }) eq))))             ≡⟨ cong (≡⇒≃ ∘ proj₁ ∘ Σ-≡,≡←≡) $
-                                                                          cong-∘ proj₁ (λ { (x , (y , z)) → (x , y) , z }) _ ⟩
-      ≡⇒≃ (proj₁ (Σ-≡,≡←≡ (cong (λ { (x , (y , z)) → x , y }) eq)))  ≡⟨ cong ≡⇒≃ $ proj₁-Σ-≡,≡←≡ _ ⟩
-
-      ≡⇒≃ (cong proj₁ (cong (λ { (x , (y , z)) → x , y }) eq))       ≡⟨ cong ≡⇒≃ $ cong-∘ proj₁ (λ { (x , (y , z)) → x , y }) _ ⟩
-
-      ≡⇒≃ (cong proj₁ eq)                                            ≡⟨ elim-cong _≃_ proj₁ _ ⟩∎
-
-      elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) eq          ∎
-
-      where open Assumptions ass
-
-    -- In fact, the entire from component of isomorphism-is-equality
-    -- is equal to a simple function.
-    --
-    -- The proof of this lemma is somewhat complicated. A much shorter
-    -- proof can be constructed if El (proj₁ c) (proj₁ J) is a set
-    -- (see
-    -- Structure-identity-principle.from-isomorphism-is-equality′).
-
-    from-isomorphism-is-equality :
-      ∀ ass c X Y →
-      _↔_.from (isomorphism-is-equality ass c X Y) ≡
-      elim (λ {X Y} _ → Isomorphic c X Y)
-           (λ { (_ , x , _) → Eq.id , resp-id ass (proj₁ c) x })
-    from-isomorphism-is-equality ass (a , P) (C , x , p) _ =
-      ext (elim¹
-        (λ eq → Σ-map ≡⇒≃ f (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡
-                  (cong (λ { (C , (x , p)) → (C , x) , p }) eq)))) ≡
-                elim (λ {X Y} _ → Isomorphic (a , P) X Y)
-                     (λ { (_ , x , _) → Eq.id , resp-id ass a x })
-                     eq)
-
-        (Σ-map ≡⇒≃ f (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡ {A = ∃ (El a)}
-                        (cong (λ { (C , (x , p)) → (C , x) , p })
-                              (refl (C , x , p))))))               ≡⟨ cong (Σ-map ≡⇒≃ f ∘ Σ-≡,≡←≡ ∘ proj₁ ∘ Σ-≡,≡←≡) $
-                                                                        cong-refl {A = Instance (a , P)}
-                                                                                  (λ { (C , (x , p)) → (C , x) , p }) ⟩
-         Σ-map ≡⇒≃ f (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡ {A = ∃ (El a)}
-                                        (refl ((C , x) , p)))))    ≡⟨ cong (Σ-map ≡⇒≃ f ∘ Σ-≡,≡←≡ ∘ proj₁) (Σ-≡,≡←≡-refl {A = ∃ (El a)}) ⟩
-
-         Σ-map ≡⇒≃ f (Σ-≡,≡←≡ (refl {A = ∃ (El a)} (C , x)))       ≡⟨ cong (Σ-map ≡⇒≃ f) (Σ-≡,≡←≡-refl {B = El a}) ⟩
-
-         (≡⇒≃ (refl C) , f (subst-refl (El a) x))                  ≡⟨ Σ-≡,≡→≡ ≡⇒≃-refl lemma₄ ⟩
-
-         (Eq.id , resp-id ass a x)                                 ≡⟨ sym $ elim-refl (λ {X Y} _ → Isomorphic (a , P) X Y) _ ⟩∎
-
-         elim (λ {X Y} _ → Isomorphic (a , P) X Y)
-              (λ { (_ , x , _) → Eq.id , resp-id ass a x })
-              (refl (C , x , p))                                   ∎))
-
-      where
-      open Assumptions ass
-
-      f : ∀ {D} {y : El a D} {eq : C ≡ D} →
-          subst (El a) eq x ≡ y →
-          resp a (≡⇒≃ eq) x ≡ y
-      f {y = y} {eq} eq′ =
-        _↔_.from (≡⇒↝ _ $ cong (λ z → z ≡ y) $
+  -- The type of (lifted) isomorphisms between two instances of a
+  -- structure is equal to the type of equalities between the same
+  -- instances (assuming univalence).
+  --
+  -- In short, isomorphism is equal to equality.
+
+  isomorphic≡≡ :
+    Assumptions →
+    ∀ c {X Y} → ↑ (# 2) (Isomorphic c X Y) ≡ (X ≡ Y)
+  isomorphic≡≡ ass c {X} {Y} =
+    ≃⇒≡ univ₂ $ ↔⇒≃ (
+      ↑ _ (Isomorphic c X Y)  ↝⟨ ↑↔ ⟩
+      Isomorphic c X Y        ↝⟨ isomorphism-is-equality ass c X Y ⟩□
+      (X ≡ Y)                 □)
+    where open Assumptions ass
+
+  -- The "first part" of the from component of
+  -- isomorphism-is-equality is equal to a simple function.
+
+  proj₁-from-isomorphism-is-equality :
+    ∀ ass c X Y →
+    proj₁ ∘ _↔_.from (isomorphism-is-equality ass c X Y) ≡
+    elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id)
+  proj₁-from-isomorphism-is-equality ass _ _ _ = ext λ eq →
+
+    ≡⇒≃ (proj₁ (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡
+      (cong (λ { (x , (y , z)) → (x , y) , z }) eq)))))            ≡⟨ cong (≡⇒≃ ∘ proj₁ ∘ Σ-≡,≡←≡) $ proj₁-Σ-≡,≡←≡ _ ⟩
+
+    ≡⇒≃ (proj₁ (Σ-≡,≡←≡ (cong proj₁
+      (cong (λ { (x , (y , z)) → (x , y) , z }) eq))))             ≡⟨ cong (≡⇒≃ ∘ proj₁ ∘ Σ-≡,≡←≡) $
+                                                                        cong-∘ proj₁ (λ { (x , (y , z)) → (x , y) , z }) _ ⟩
+    ≡⇒≃ (proj₁ (Σ-≡,≡←≡ (cong (λ { (x , (y , z)) → x , y }) eq)))  ≡⟨ cong ≡⇒≃ $ proj₁-Σ-≡,≡←≡ _ ⟩
+
+    ≡⇒≃ (cong proj₁ (cong (λ { (x , (y , z)) → x , y }) eq))       ≡⟨ cong ≡⇒≃ $ cong-∘ proj₁ (λ { (x , (y , z)) → x , y }) _ ⟩
+
+    ≡⇒≃ (cong proj₁ eq)                                            ≡⟨ elim-cong _≃_ proj₁ _ ⟩∎
+
+    elim (λ {X Y} _ → proj₁ X ≃ proj₁ Y) (λ _ → Eq.id) eq          ∎
+
+    where open Assumptions ass
+
+  -- In fact, the entire from component of isomorphism-is-equality
+  -- is equal to a simple function.
+  --
+  -- The proof of this lemma is somewhat complicated. A much shorter
+  -- proof can be constructed if El (proj₁ c) (proj₁ J) is a set
+  -- (see
+  -- Structure-identity-principle.from-isomorphism-is-equality′).
+
+  from-isomorphism-is-equality :
+    ∀ ass c X Y →
+    _↔_.from (isomorphism-is-equality ass c X Y) ≡
+    elim (λ {X Y} _ → Isomorphic c X Y)
+         (λ { (_ , x , _) → Eq.id , resp-id ass (proj₁ c) x })
+  from-isomorphism-is-equality ass (a , P) (C , x , p) _ =
+    ext (elim¹
+      (λ eq → Σ-map ≡⇒≃ f (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡
+                (cong (λ { (C , (x , p)) → (C , x) , p }) eq)))) ≡
+              elim (λ {X Y} _ → Isomorphic (a , P) X Y)
+                   (λ { (_ , x , _) → Eq.id , resp-id ass a x })
+                   eq)
+
+      (Σ-map ≡⇒≃ f (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡ {A = ∃ (El a)}
+                      (cong (λ { (C , (x , p)) → (C , x) , p })
+                            (refl (C , x , p))))))               ≡⟨ cong (Σ-map ≡⇒≃ f ∘ Σ-≡,≡←≡ ∘ proj₁ ∘ Σ-≡,≡←≡) $
+                                                                      cong-refl {A = Instance (a , P)}
+                                                                                (λ { (C , (x , p)) → (C , x) , p }) ⟩
+       Σ-map ≡⇒≃ f (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡ {A = ∃ (El a)}
+                                      (refl ((C , x) , p)))))    ≡⟨ cong (Σ-map ≡⇒≃ f ∘ Σ-≡,≡←≡ ∘ proj₁) (Σ-≡,≡←≡-refl {A = ∃ (El a)}) ⟩
+
+       Σ-map ≡⇒≃ f (Σ-≡,≡←≡ (refl {A = ∃ (El a)} (C , x)))       ≡⟨ cong (Σ-map ≡⇒≃ f) (Σ-≡,≡←≡-refl {B = El a}) ⟩
+
+       (≡⇒≃ (refl C) , f (subst-refl (El a) x))                  ≡⟨ Σ-≡,≡→≡ ≡⇒≃-refl lemma₄ ⟩
+
+       (Eq.id , resp-id ass a x)                                 ≡⟨ sym $ elim-refl (λ {X Y} _ → Isomorphic (a , P) X Y) _ ⟩∎
+
+       elim (λ {X Y} _ → Isomorphic (a , P) X Y)
+            (λ { (_ , x , _) → Eq.id , resp-id ass a x })
+            (refl (C , x , p))                                   ∎))
+
+    where
+    open Assumptions ass
+
+    f : ∀ {D} {y : El a D} {eq : C ≡ D} →
+        subst (El a) eq x ≡ y →
+        resp a (≡⇒≃ eq) x ≡ y
+    f {y = y} {eq} eq′ =
+      _↔_.from (≡⇒↝ _ $ cong (λ z → z ≡ y) $
+                  transport-theorem (El a) (resp a) (resp-id ass a)
+                                    univ₁ (≡⇒≃ eq) x)
+         (_↔_.to (≡⇒↝ _ $ sym $ cong (λ eq → subst (El a) eq x ≡ y)
+                    (_≃_.left-inverse-of (≡≃≃ univ₁) eq)) eq′)
+
+    lemma₁ : ∀ {ℓ} {A B C : Set ℓ} {x} (eq₁ : B ≡ A) (eq₂ : C ≡ B) →
+             _↔_.from (≡⇒↝ _ eq₂) (_↔_.to (≡⇒↝ _ (sym eq₁)) x) ≡
+             _↔_.to (≡⇒↝ _ (sym (trans eq₂ eq₁))) x
+    lemma₁ {x = x} eq₁ eq₂ =
+      _↔_.from (≡⇒↝ _ eq₂) (_↔_.to (≡⇒↝ _ (sym eq₁)) x)      ≡⟨ sym $ cong (λ f → f (_↔_.to (≡⇒↝ _ (sym eq₁)) x)) $ ≡⇒↝-sym bijection ⟩
+      _↔_.to (≡⇒↝ _ (sym eq₂)) (_↔_.to (≡⇒↝ _ (sym eq₁)) x)  ≡⟨ sym $ cong (λ f → f x) $ ≡⇒↝-trans bijection ⟩
+      _↔_.to (≡⇒↝ _ (trans (sym eq₁) (sym eq₂))) x           ≡⟨ sym $ cong (λ eq → _↔_.to (≡⇒↝ _ eq) x) $ sym-trans _ _ ⟩∎
+      _↔_.to (≡⇒↝ _ (sym (trans eq₂ eq₁))) x                 ∎
+
+    lemma₂ : ∀ {a} {A : Set a} {x y z : A}
+             (x≡y : x ≡ y) (y≡z : y ≡ z) →
+             _↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) (sym x≡y))) y≡z ≡
+             trans x≡y y≡z
+    lemma₂ {y = y} {z} x≡y y≡z = elim₁
+      (λ x≡y → _↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) (sym x≡y))) y≡z ≡
+               trans x≡y y≡z)
+      (_↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) (sym (refl y)))) y≡z  ≡⟨ cong (λ eq → _↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) eq)) y≡z) sym-refl ⟩
+       _↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) (refl y))) y≡z        ≡⟨ cong (λ eq → _↔_.to (≡⇒↝ _ eq) y≡z) $ cong-refl (λ x → x ≡ z) ⟩
+       _↔_.to (≡⇒↝ _ (refl (y ≡ z))) y≡z                       ≡⟨ cong (λ f → _↔_.to f y≡z) ≡⇒↝-refl ⟩
+       y≡z                                                     ≡⟨ sym $ trans-reflˡ _ ⟩∎
+       trans (refl y) y≡z                                      ∎)
+      x≡y
+
+    lemma₃ :
+      sym (trans (cong (λ z → z ≡ x) $
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple.agda 343
-                                      univ₁ (≡⇒≃ eq) x)
-           (_↔_.to (≡⇒↝ _ $ sym $ cong (λ eq → subst (El a) eq x ≡ y)
-                      (_≃_.left-inverse-of (≡≃≃ univ₁) eq)) eq′)
-
-      lemma₁ : ∀ {ℓ} {A B C : Set ℓ} {x} (eq₁ : B ≡ A) (eq₂ : C ≡ B) →
-               _↔_.from (≡⇒↝ _ eq₂) (_↔_.to (≡⇒↝ _ (sym eq₁)) x) ≡
-               _↔_.to (≡⇒↝ _ (sym (trans eq₂ eq₁))) x
-      lemma₁ {x = x} eq₁ eq₂ =
-        _↔_.from (≡⇒↝ _ eq₂) (_↔_.to (≡⇒↝ _ (sym eq₁)) x)      ≡⟨ sym $ cong (λ f → f (_↔_.to (≡⇒↝ _ (sym eq₁)) x)) $ ≡⇒↝-sym bijection ⟩
-        _↔_.to (≡⇒↝ _ (sym eq₂)) (_↔_.to (≡⇒↝ _ (sym eq₁)) x)  ≡⟨ sym $ cong (λ f → f x) $ ≡⇒↝-trans bijection ⟩
-        _↔_.to (≡⇒↝ _ (trans (sym eq₁) (sym eq₂))) x           ≡⟨ sym $ cong (λ eq → _↔_.to (≡⇒↝ _ eq) x) $ sym-trans _ _ ⟩∎
-        _↔_.to (≡⇒↝ _ (sym (trans eq₂ eq₁))) x                 ∎
-
-      lemma₂ : ∀ {a} {A : Set a} {x y z : A}
-               (x≡y : x ≡ y) (y≡z : y ≡ z) →
-               _↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) (sym x≡y))) y≡z ≡
-               trans x≡y y≡z
-      lemma₂ {y = y} {z} x≡y y≡z = elim₁
-        (λ x≡y → _↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) (sym x≡y))) y≡z ≡
-                 trans x≡y y≡z)
-        (_↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) (sym (refl y)))) y≡z  ≡⟨ cong (λ eq → _↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) eq)) y≡z) sym-refl ⟩
-         _↔_.to (≡⇒↝ _ (cong (λ x → x ≡ z) (refl y))) y≡z        ≡⟨ cong (λ eq → _↔_.to (≡⇒↝ _ eq) y≡z) $ cong-refl (λ x → x ≡ z) ⟩
-         _↔_.to (≡⇒↝ _ (refl (y ≡ z))) y≡z                       ≡⟨ cong (λ f → _↔_.to f y≡z) ≡⇒↝-refl ⟩
-         y≡z                                                     ≡⟨ sym $ trans-reflˡ _ ⟩∎
-         trans (refl y) y≡z                                      ∎)
-        x≡y
-
-      lemma₃ :
-        sym (trans (cong (λ z → z ≡ x) $
-                      transport-theorem (El a) (resp a) (resp-id ass a)
+                                      univ₁ (≡⇒≃ (refl C)) x)
+             (cong (λ eq → subst (El a) eq x ≡ x)
+                (_≃_.left-inverse-of (≡≃≃ univ₁) (refl C)))) ≡
+      cong (λ z → z ≡ x) (sym $
+        trans (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
+                 (resp-id ass a x))
+          (sym $ subst-refl (El a) x))
+    lemma₃ =
+      sym (trans (cong (λ z → z ≡ x) _)
+             (cong (λ eq → subst (El a) eq x ≡ x) _))           ≡⟨ cong (λ eq → sym (trans (cong (λ z → z ≡ x)
+                                                                                              (transport-theorem (El a) (resp a) (resp-id ass a)
+                                                                                                                 univ₁ (≡⇒≃ (refl C)) x))
+                                                                                       eq)) $ sym $
+                                                                        cong-∘ (λ z → z ≡ x) (λ eq → subst (El a) eq x) _ ⟩
+      sym (trans (cong (λ z → z ≡ x) _)
+             (cong (λ z → z ≡ x)
+                (cong (λ eq → subst (El a) eq x) _)))           ≡⟨ cong sym $ sym $ cong-trans (λ z → z ≡ x) _ _ ⟩
+
+      sym (cong (λ z → z ≡ x)
+        (trans _ (cong (λ eq → subst (El a) eq x) _)))          ≡⟨ sym $ cong-sym (λ z → z ≡ x) _ ⟩
+
+      cong (λ z → z ≡ x) (sym $
+        trans (transport-theorem (El a) (resp a)
+                 (resp-id ass a) univ₁ (≡⇒≃ (refl C)) x)
+          (cong (λ eq → subst (El a) eq x) _))                  ≡⟨ cong (λ eq → cong (λ z → z ≡ x) (sym $
+                                                                                  trans eq (cong (λ eq → subst (El a) eq x)
+                                                                                              (_≃_.left-inverse-of (≡≃≃ univ₁) (refl C)))))
+                                                                        (transport-theorem-≡⇒≃-refl (El a) (resp a) (resp-id ass a) univ₁ _) ⟩
+      cong (λ z → z ≡ x) (sym $
+        trans (trans (trans (trans (cong (λ eq → resp a eq x)
+                                      ≡⇒≃-refl)
+                               (resp-id ass a x))
+                        (sym $ subst-refl (El a) x))
+                 (sym $ cong (λ eq → subst (El a) eq x)
+                          (_≃_.left-inverse-of
+                             (≡≃≃ univ₁) (refl C))))
+          (cong (λ eq → subst (El a) eq x)
+                (_≃_.left-inverse-of (≡≃≃ univ₁) (refl C))))    ≡⟨ cong (cong (λ z → z ≡ x) ∘ sym) $
+                                                                        trans-[trans-sym] _ _ ⟩∎
+      cong (λ z → z ≡ x) (sym $
+         trans (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
+                  (resp-id ass a x))
+           (sym $ subst-refl (El a) x))                         ∎
+
+    lemma₄ : subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
+                   (f (subst-refl (El a) x)) ≡
+             resp-id ass a x
+    lemma₄ =
+      subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
+        (f (subst-refl (El a) x))                                      ≡⟨ cong (subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl) $ lemma₁ _ _ ⟩
+
+      subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
+        (_↔_.to (≡⇒↝ _ (sym (trans (cong (λ z → z ≡ x) $
+                                      transport-theorem (El a)
+                                        (resp a) (resp-id ass a)
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple.agda 399
-               (cong (λ eq → subst (El a) eq x ≡ x)
-                  (_≃_.left-inverse-of (≡≃≃ univ₁) (refl C)))) ≡
-        cong (λ z → z ≡ x) (sym $
-          trans (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
-                   (resp-id ass a x))
-            (sym $ subst-refl (El a) x))
-      lemma₃ =
-        sym (trans (cong (λ z → z ≡ x) _)
-               (cong (λ eq → subst (El a) eq x ≡ x) _))           ≡⟨ cong (λ eq → sym (trans (cong (λ z → z ≡ x)
-                                                                                                (transport-theorem (El a) (resp a) (resp-id ass a)
-                                                                                                                   univ₁ (≡⇒≃ (refl C)) x))
-                                                                                         eq)) $ sym $
-                                                                          cong-∘ (λ z → z ≡ x) (λ eq → subst (El a) eq x) _ ⟩
-        sym (trans (cong (λ z → z ≡ x) _)
-               (cong (λ z → z ≡ x)
-                  (cong (λ eq → subst (El a) eq x) _)))           ≡⟨ cong sym $ sym $ cong-trans (λ z → z ≡ x) _ _ ⟩
-
-        sym (cong (λ z → z ≡ x)
-          (trans _ (cong (λ eq → subst (El a) eq x) _)))          ≡⟨ sym $ cong-sym (λ z → z ≡ x) _ ⟩
-
-        cong (λ z → z ≡ x) (sym $
-          trans (transport-theorem (El a) (resp a)
-                   (resp-id ass a) univ₁ (≡⇒≃ (refl C)) x)
-            (cong (λ eq → subst (El a) eq x) _))                  ≡⟨ cong (λ eq → cong (λ z → z ≡ x) (sym $
-                                                                                    trans eq (cong (λ eq → subst (El a) eq x)
-                                                                                                (_≃_.left-inverse-of (≡≃≃ univ₁) (refl C)))))
-                                                                          (transport-theorem-≡⇒≃-refl (El a) (resp a) (resp-id ass a) univ₁ _) ⟩
-        cong (λ z → z ≡ x) (sym $
-          trans (trans (trans (trans (cong (λ eq → resp a eq x)
-                                        ≡⇒≃-refl)
-                                 (resp-id ass a x))
-                          (sym $ subst-refl (El a) x))
-                   (sym $ cong (λ eq → subst (El a) eq x)
-                            (_≃_.left-inverse-of
-                               (≡≃≃ univ₁) (refl C))))
-            (cong (λ eq → subst (El a) eq x)
-                  (_≃_.left-inverse-of (≡≃≃ univ₁) (refl C))))    ≡⟨ cong (cong (λ z → z ≡ x) ∘ sym) $
-                                                                          trans-[trans-sym] _ _ ⟩∎
-        cong (λ z → z ≡ x) (sym $
-           trans (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
-                    (resp-id ass a x))
-             (sym $ subst-refl (El a) x))                         ∎
-
-      lemma₄ : subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
-                     (f (subst-refl (El a) x)) ≡
-               resp-id ass a x
-      lemma₄ =
-        subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
-          (f (subst-refl (El a) x))                                      ≡⟨ cong (subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl) $ lemma₁ _ _ ⟩
-
-        subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
-          (_↔_.to (≡⇒↝ _ (sym (trans (cong (λ z → z ≡ x) $
-                                        transport-theorem (El a)
-                                          (resp a) (resp-id ass a)
-                                          univ₁ (≡⇒≃ (refl C)) x)
-                                 (cong (λ eq → subst (El a) eq x ≡ x)
-                                    (_≃_.left-inverse-of
-                                       (≡≃≃ univ₁) (refl C))))))
-                  (subst-refl (El a) x))                                 ≡⟨ cong (λ eq → subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
-                                                                                           (_↔_.to (≡⇒↝ _ eq) (subst-refl (El a) x)))
-                                                                                 lemma₃ ⟩
-        subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
-          (_↔_.to (≡⇒↝ _
-                     (cong (λ z → z ≡ x) $ sym
-                        (trans (trans (cong (λ eq → resp a eq x)
-                                         ≡⇒≃-refl)
-                                  (resp-id ass a x))
-                           (sym $ subst-refl (El a) x))))
-                  (subst-refl (El a) x))                                 ≡⟨ cong (subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl) $ lemma₂ _ _ ⟩
-
-        subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
-          (trans (trans (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
-                           (resp-id ass a x))
-                    (sym $ subst-refl (El a) x))
-             (subst-refl (El a) x))                                      ≡⟨ cong (λ eq → subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl eq)
-                                                                                 (trans-[trans-sym] _ _) ⟩
-        subst (λ eq → resp a eq x ≡ x) ≡⇒≃-refl
-          (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
-             (resp-id ass a x))                                          ≡⟨ subst-∘ (λ z → z ≡ x) (λ eq → resp a eq x) _ ⟩
-
-        subst (λ z → z ≡ x)
-          (cong (λ eq → resp a eq x) ≡⇒≃-refl)
-          (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
-             (resp-id ass a x))                                          ≡⟨ cong (λ eq → subst (λ z → z ≡ x) eq
-                                                                                           (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
-                                                                                              (resp-id ass a x))) $
-                                                                                 sym $ sym-sym _ ⟩
-        subst (λ z → z ≡ x)
-          (sym $ sym $ cong (λ eq → resp a eq x) ≡⇒≃-refl)
-          (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
-             (resp-id ass a x))                                          ≡⟨ subst-trans (sym $ cong (λ eq → resp a eq x) ≡⇒≃-refl) ⟩
-
-        trans (sym $ cong (λ eq → resp a eq x) ≡⇒≃-refl)
-          (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
-             (resp-id ass a x))                                          ≡⟨ sym $ trans-assoc _ _ _ ⟩
-
-        trans (trans (sym $ cong (λ eq → resp a eq x) ≡⇒≃-refl)
-                     (cong (λ eq → resp a eq x) ≡⇒≃-refl))
-          (resp-id ass a x)                                              ≡⟨ cong (λ eq → trans eq _) $ trans-symˡ _ ⟩
-
-        trans (refl (resp a Eq.id x)) (resp-id ass a x)                  ≡⟨ trans-reflˡ _ ⟩∎
-
-        resp-id ass a x                                                  ∎
+                               (cong (λ eq → subst (El a) eq x ≡ x)
+                                  (_≃_.left-inverse-of
+                                     (≡≃≃ univ₁) (refl C))))))
+                (subst-refl (El a) x))                                 ≡⟨ cong (λ eq → subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
+                                                                                         (_↔_.to (≡⇒↝ _ eq) (subst-refl (El a) x)))
+                                                                               lemma₃ ⟩
+      subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
+        (_↔_.to (≡⇒↝ _
+                   (cong (λ z → z ≡ x) $ sym
+                      (trans (trans (cong (λ eq → resp a eq x)
+                                       ≡⇒≃-refl)
+                                (resp-id ass a x))
+                         (sym $ subst-refl (El a) x))))
+                (subst-refl (El a) x))                                 ≡⟨ cong (subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl) $ lemma₂ _ _ ⟩
+
+      subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl
+        (trans (trans (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
+                         (resp-id ass a x))
+                  (sym $ subst-refl (El a) x))
+           (subst-refl (El a) x))                                      ≡⟨ cong (λ eq → subst (λ eq → Is-isomorphism a eq x x) ≡⇒≃-refl eq)
+                                                                               (trans-[trans-sym] _ _) ⟩
+      subst (λ eq → resp a eq x ≡ x) ≡⇒≃-refl
+        (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
+           (resp-id ass a x))                                          ≡⟨ subst-∘ (λ z → z ≡ x) (λ eq → resp a eq x) _ ⟩
+
+      subst (λ z → z ≡ x)
+        (cong (λ eq → resp a eq x) ≡⇒≃-refl)
+        (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
+           (resp-id ass a x))                                          ≡⟨ cong (λ eq → subst (λ z → z ≡ x) eq
+                                                                                         (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
+                                                                                            (resp-id ass a x))) $
+                                                                               sym $ sym-sym _ ⟩
+      subst (λ z → z ≡ x)
+        (sym $ sym $ cong (λ eq → resp a eq x) ≡⇒≃-refl)
+        (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
+           (resp-id ass a x))                                          ≡⟨ subst-trans (sym $ cong (λ eq → resp a eq x) ≡⇒≃-refl) ⟩
+
+      trans (sym $ cong (λ eq → resp a eq x) ≡⇒≃-refl)
+        (trans (cong (λ eq → resp a eq x) ≡⇒≃-refl)
+           (resp-id ass a x))                                          ≡⟨ sym $ trans-assoc _ _ _ ⟩
+
+      trans (trans (sym $ cong (λ eq → resp a eq x) ≡⇒≃-refl)
+                   (cong (λ eq → resp a eq x) ≡⇒≃-refl))
+        (resp-id ass a x)                                              ≡⟨ cong (λ eq → trans eq _) $ trans-symˡ _ ⟩
+
+      trans (refl (resp a Eq.id x)) (resp-id ass a x)                  ≡⟨ trans-reflˡ _ ⟩∎
+
+      resp-id ass a x                                                  ∎
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple.agda 490
-abstract
-
-  -- The cast function respects identities (assuming extensionality).
-
-  cast-id : Extensionality (# 1) (# 1) →
-            ∀ a {B} → cast a (Logical-equivalence.id {A = B}) ≡
-                      Logical-equivalence.id
-  cast-id ext id      = refl _
-  cast-id ext set     = refl _
-  cast-id ext (k A)   = refl _
-  cast-id ext (a ⇾ b) = cong₂ →-cong-⇔ (cast-id ext a) (cast-id ext b)
-  cast-id ext (a ⊗ b) = cong₂ _×-cong_ (cast-id ext a) (cast-id ext b)
-  cast-id ext (a ⊕ b) =
-    cast a Logical-equivalence.id ⊎-cong cast b Logical-equivalence.id  ≡⟨ cong₂ _⊎-cong_ (cast-id ext a) (cast-id ext b) ⟩
-    Logical-equivalence.id ⊎-cong Logical-equivalence.id                ≡⟨ cong₂ (λ f g → record { to = f; from = g })
-                                                                                 (ext [ refl ∘ inj₁ , refl ∘ inj₂ ])
-                                                                                 (ext [ refl ∘ inj₁ , refl ∘ inj₂ ]) ⟩∎
-    Logical-equivalence.id                                              ∎
-
-  resp-id : Extensionality (# 1) (# 1) →
-            ∀ a {B} x → resp a (Eq.id {A = B}) x ≡ x
-  resp-id ext a x = cong (λ eq → _⇔_.to eq x) $ cast-id ext a
+-- The cast function respects identities (assuming extensionality).
+
+cast-id : Extensionality (# 1) (# 1) →
+          ∀ a {B} → cast a (Logical-equivalence.id {A = B}) ≡
+                    Logical-equivalence.id
+cast-id ext id      = refl _
+cast-id ext set     = refl _
+cast-id ext (k A)   = refl _
+cast-id ext (a ⇾ b) = cong₂ →-cong-⇔ (cast-id ext a) (cast-id ext b)
+cast-id ext (a ⊗ b) = cong₂ _×-cong_ (cast-id ext a) (cast-id ext b)
+cast-id ext (a ⊕ b) =
+  cast a Logical-equivalence.id ⊎-cong cast b Logical-equivalence.id  ≡⟨ cong₂ _⊎-cong_ (cast-id ext a) (cast-id ext b) ⟩
+  Logical-equivalence.id ⊎-cong Logical-equivalence.id                ≡⟨ cong₂ (λ f g → record { to = f; from = g })
+                                                                               (ext [ refl ∘ inj₁ , refl ∘ inj₂ ])
+                                                                               (ext [ refl ∘ inj₁ , refl ∘ inj₂ ]) ⟩∎
+  Logical-equivalence.id                                              ∎
+
+resp-id : Extensionality (# 1) (# 1) →
+          ∀ a {B} x → resp a (Eq.id {A = B}) x ≡ x
+resp-id ext a x = cong (λ eq → _⇔_.to eq x) $ cast-id ext a
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple.agda 575
-  abstract
-
-    -- The projection _≃_.logical-equivalence is homomorphic with
-    -- respect to cast a/cst a.
-
-    casts-related : ∀ a →
-                    cast a (_≃_.logical-equivalence B≃C) ≡
-                    _≃_.logical-equivalence (cst a)
-    casts-related id      = refl _
-    casts-related set     = refl _
-    casts-related (k A)   = refl _
-    casts-related (a ⇾ b) = cong₂ →-cong-⇔ (casts-related a)
-                                           (casts-related b)
-    casts-related (a ⊗ b) = cong₂ _×-cong_ (casts-related a)
-                                           (casts-related b)
-    casts-related (a ⊕ b) = cong₂ _⊎-cong_ (casts-related a)
-                                           (casts-related b)
-
-    to∘from : ∀ x → _⇔_.to (cast a B⇔C) (_⇔_.from (cast a B⇔C) x) ≡ x
-    to∘from x =
-      _⇔_.to (cast a B⇔C) (_⇔_.from (cast a B⇔C) x)  ≡⟨ cong₂ (λ f g → f (g x)) (cong _⇔_.to $ casts-related a)
-                                                                                (cong _⇔_.from $ casts-related a) ⟩
-      _≃_.to (cst a) (_≃_.from (cst a) x)            ≡⟨ _≃_.right-inverse-of (cst a) x ⟩∎
-      x                                              ∎
-
-    from∘to : ∀ x → _⇔_.from (cast a B⇔C) (_⇔_.to (cast a B⇔C) x) ≡ x
-    from∘to x =
-      _⇔_.from (cast a B⇔C) (_⇔_.to (cast a B⇔C) x)  ≡⟨ cong₂ (λ f g → f (g x)) (cong _⇔_.from $ casts-related a)
-                                                                                (cong _⇔_.to $ casts-related a) ⟩
-      _≃_.from (cst a) (_≃_.to (cst a) x)            ≡⟨ _≃_.left-inverse-of (cst a) x ⟩∎
-      x                                              ∎
+  -- The projection _≃_.logical-equivalence is homomorphic with
+  -- respect to cast a/cst a.
+
+  casts-related : ∀ a →
+                  cast a (_≃_.logical-equivalence B≃C) ≡
+                  _≃_.logical-equivalence (cst a)
+  casts-related id      = refl _
+  casts-related set     = refl _
+  casts-related (k A)   = refl _
+  casts-related (a ⇾ b) = cong₂ →-cong-⇔ (casts-related a)
+                                         (casts-related b)
+  casts-related (a ⊗ b) = cong₂ _×-cong_ (casts-related a)
+                                         (casts-related b)
+  casts-related (a ⊕ b) = cong₂ _⊎-cong_ (casts-related a)
+                                         (casts-related b)
+
+  to∘from : ∀ x → _⇔_.to (cast a B⇔C) (_⇔_.from (cast a B⇔C) x) ≡ x
+  to∘from x =
+    _⇔_.to (cast a B⇔C) (_⇔_.from (cast a B⇔C) x)  ≡⟨ cong₂ (λ f g → f (g x)) (cong _⇔_.to $ casts-related a)
+                                                                              (cong _⇔_.from $ casts-related a) ⟩
+    _≃_.to (cst a) (_≃_.from (cst a) x)            ≡⟨ _≃_.right-inverse-of (cst a) x ⟩∎
+    x                                              ∎
+
+  from∘to : ∀ x → _⇔_.from (cast a B⇔C) (_⇔_.to (cast a B⇔C) x) ≡ x
+  from∘to x =
+    _⇔_.from (cast a B⇔C) (_⇔_.to (cast a B⇔C) x)  ≡⟨ cong₂ (λ f g → f (g x)) (cong _⇔_.from $ casts-related a)
+                                                                              (cong _⇔_.to $ casts-related a) ⟩
+    _≃_.from (cst a) (_≃_.to (cst a) x)            ≡⟨ _≃_.left-inverse-of (cst a) x ⟩∎
+    x                                              ∎
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple.agda 627
-abstract
-
-  -- The two definitions of "being an isomorphism" are "isomorphic"
-  -- (in bijective correspondence), assuming univalence.
-
-  is-isomorphism-isomorphic :
-    Assumptions →
-    ∀ a {B C x y} (eq : B ≃ C) →
-    Is-isomorphism a eq x y ↔ Is-isomorphism′ a eq x y
-
-  is-isomorphism-isomorphic ass id {x = x} {y} eq =
-
-    (_≃_.to eq x ≡ y)  □
-
-  is-isomorphism-isomorphic ass set {x = X} {Y} eq =
-
-    (X ≡ Y)      ↔⟨ ≡≃≃ univ ⟩
-
-    (X ≃ Y)      ↝⟨ inverse ↑↔ ⟩□
-
-    ↑ _ (X ≃ Y)  □
-
-    where open Assumptions ass
-
-  is-isomorphism-isomorphic ass (k A) {x = x} {y} eq =
-
-    (x ≡ y) □
-
-  is-isomorphism-isomorphic ass (a ⇾ b) {x = f} {g} eq =
-
-    (resp b eq ∘ f ∘ resp⁻¹ a eq ≡ g)                  ↝⟨ ∘from≡↔≡∘to ext₁ (cast≃ ext₁ a eq) ⟩
-
-    (resp b eq ∘ f ≡ g ∘ resp a eq)                    ↔⟨ inverse $ extensionality-isomorphism ext₁ ⟩
-
-    (∀ x → resp b eq (f x) ≡ g (resp a eq x))          ↔⟨ ∀-preserves ext₁ (λ x → ↔⇒≃ $
-                                                            ∀-intro ext₁ (λ y _ → resp b eq (f x) ≡ g y)) ⟩
-    (∀ x y → resp a eq x ≡ y → resp b eq (f x) ≡ g y)  ↔⟨ ∀-preserves ext₁ (λ _ → ∀-preserves ext₁ λ _ → ↔⇒≃ $
-                                                            →-cong ext₁ (is-isomorphism-isomorphic ass a eq)
-                                                                        (is-isomorphism-isomorphic ass b eq)) ⟩□
-    (∀ x y → Is-isomorphism′ a eq x y →
-             Is-isomorphism′ b eq (f x) (g y))         □
-
-    where open Assumptions ass
-
-  is-isomorphism-isomorphic ass (a ⊗ b) {x = x , u} {y , v} eq =
-
-    ((resp a eq x , resp b eq u) ≡ (y , v))              ↝⟨ inverse ≡×≡↔≡ ⟩
-
-    (resp a eq x ≡ y × resp b eq u ≡ v)                  ↝⟨ is-isomorphism-isomorphic ass a eq ×-cong
-                                                            is-isomorphism-isomorphic ass b eq ⟩□
-    Is-isomorphism′ a eq x y × Is-isomorphism′ b eq u v  □
-
-    where open Assumptions ass
-
-  is-isomorphism-isomorphic ass (a ⊕ b) {x = inj₁ x} {inj₁ y} eq =
-
-    (inj₁ (resp a eq x) ≡ inj₁ y)  ↝⟨ inverse ≡↔inj₁≡inj₁ ⟩
-
-    (resp a eq x ≡ y)              ↝⟨ is-isomorphism-isomorphic ass a eq ⟩□
-
-    Is-isomorphism′ a eq x y       □
-
-    where open Assumptions ass
-
-  is-isomorphism-isomorphic ass (a ⊕ b) {x = inj₂ x} {inj₂ y} eq =
-
-    (inj₂ (resp b eq x) ≡ inj₂ y)  ↝⟨ inverse ≡↔inj₂≡inj₂ ⟩
-
-    (resp b eq x ≡ y)              ↝⟨ is-isomorphism-isomorphic ass b eq ⟩□
-
-    Is-isomorphism′ b eq x y       □
-
-    where open Assumptions ass
-
-  is-isomorphism-isomorphic ass (a ⊕ b) {x = inj₁ x} {inj₂ y} eq =
-
-    (inj₁ _ ≡ inj₂ _)  ↝⟨ inverse $ ⊥↔uninhabited ⊎.inj₁≢inj₂ ⟩□
-
-    ⊥                  □
-
-  is-isomorphism-isomorphic ass (a ⊕ b) {x = inj₂ x} {inj₁ y} eq =
-
-    (inj₂ _ ≡ inj₁ _)  ↝⟨ inverse $ ⊥↔uninhabited (⊎.inj₁≢inj₂ ∘ sym) ⟩□
-
-    ⊥                  □
-
-  -- The two definitions of isomorphism are "isomorphic" (in bijective
-  -- correspondence), assuming univalence.
-
-  isomorphic-isomorphic :
-    Assumptions →
-    ∀ c X Y →
-    Isomorphic c X Y ↔ Isomorphic′ c X Y
-  isomorphic-isomorphic ass (a , _) (C , x , _) (D , y , _) =
-    Σ (C ≃ D) (λ eq → Is-isomorphism  a eq x y)  ↝⟨ ∃-cong (λ eq → is-isomorphism-isomorphic ass a eq) ⟩
-    Σ (C ≃ D) (λ eq → Is-isomorphism′ a eq x y)  □
+-- The two definitions of "being an isomorphism" are "isomorphic"
+-- (in bijective correspondence), assuming univalence.
+
+is-isomorphism-isomorphic :
+  Assumptions →
+  ∀ a {B C x y} (eq : B ≃ C) →
+  Is-isomorphism a eq x y ↔ Is-isomorphism′ a eq x y
+
+is-isomorphism-isomorphic ass id {x = x} {y} eq =
+
+  (_≃_.to eq x ≡ y)  □
+
+is-isomorphism-isomorphic ass set {x = X} {Y} eq =
+
+  (X ≡ Y)      ↔⟨ ≡≃≃ univ ⟩
+
+  (X ≃ Y)      ↝⟨ inverse ↑↔ ⟩□
+
+  ↑ _ (X ≃ Y)  □
+
+  where open Assumptions ass
+
+is-isomorphism-isomorphic ass (k A) {x = x} {y} eq =
+
+  (x ≡ y) □
+
+is-isomorphism-isomorphic ass (a ⇾ b) {x = f} {g} eq =
+
+  (resp b eq ∘ f ∘ resp⁻¹ a eq ≡ g)                  ↝⟨ ∘from≡↔≡∘to ext₁ (cast≃ ext₁ a eq) ⟩
+
+  (resp b eq ∘ f ≡ g ∘ resp a eq)                    ↔⟨ inverse $ extensionality-isomorphism ext₁ ⟩
+
+  (∀ x → resp b eq (f x) ≡ g (resp a eq x))          ↔⟨ ∀-preserves ext₁ (λ x → ↔⇒≃ $
+                                                          ∀-intro ext₁ (λ y _ → resp b eq (f x) ≡ g y)) ⟩
+  (∀ x y → resp a eq x ≡ y → resp b eq (f x) ≡ g y)  ↔⟨ ∀-preserves ext₁ (λ _ → ∀-preserves ext₁ λ _ → ↔⇒≃ $
+                                                          →-cong ext₁ (is-isomorphism-isomorphic ass a eq)
+                                                                      (is-isomorphism-isomorphic ass b eq)) ⟩□
+  (∀ x y → Is-isomorphism′ a eq x y →
+           Is-isomorphism′ b eq (f x) (g y))         □
+
+  where open Assumptions ass
+
+is-isomorphism-isomorphic ass (a ⊗ b) {x = x , u} {y , v} eq =
+
+  ((resp a eq x , resp b eq u) ≡ (y , v))              ↝⟨ inverse ≡×≡↔≡ ⟩
+
+  (resp a eq x ≡ y × resp b eq u ≡ v)                  ↝⟨ is-isomorphism-isomorphic ass a eq ×-cong
+                                                          is-isomorphism-isomorphic ass b eq ⟩□
+  Is-isomorphism′ a eq x y × Is-isomorphism′ b eq u v  □
+
+  where open Assumptions ass
+
+is-isomorphism-isomorphic ass (a ⊕ b) {x = inj₁ x} {inj₁ y} eq =
+
+  (inj₁ (resp a eq x) ≡ inj₁ y)  ↝⟨ inverse ≡↔inj₁≡inj₁ ⟩
+
+  (resp a eq x ≡ y)              ↝⟨ is-isomorphism-isomorphic ass a eq ⟩□
+
+  Is-isomorphism′ a eq x y       □
+
+  where open Assumptions ass
+
+is-isomorphism-isomorphic ass (a ⊕ b) {x = inj₂ x} {inj₂ y} eq =
+
+  (inj₂ (resp b eq x) ≡ inj₂ y)  ↝⟨ inverse ≡↔inj₂≡inj₂ ⟩
+
+  (resp b eq x ≡ y)              ↝⟨ is-isomorphism-isomorphic ass b eq ⟩□
+
+  Is-isomorphism′ b eq x y       □
+
+  where open Assumptions ass
+
+is-isomorphism-isomorphic ass (a ⊕ b) {x = inj₁ x} {inj₂ y} eq =
+
+  (inj₁ _ ≡ inj₂ _)  ↝⟨ inverse $ ⊥↔uninhabited ⊎.inj₁≢inj₂ ⟩□
+
+  ⊥                  □
+
+is-isomorphism-isomorphic ass (a ⊕ b) {x = inj₂ x} {inj₁ y} eq =
+
+  (inj₂ _ ≡ inj₁ _)  ↝⟨ inverse $ ⊥↔uninhabited (⊎.inj₁≢inj₂ ∘ sym) ⟩□
+
+  ⊥                  □
+
+-- The two definitions of isomorphism are "isomorphic" (in bijective
+-- correspondence), assuming univalence.
+
+isomorphic-isomorphic :
+  Assumptions →
+  ∀ c X Y →
+  Isomorphic c X Y ↔ Isomorphic′ c X Y
+isomorphic-isomorphic ass (a , _) (C , x , _) (D , y , _) =
+  Σ (C ≃ D) (λ eq → Is-isomorphism  a eq x y)  ↝⟨ ∃-cong (λ eq → is-isomorphism-isomorphic ass a eq) ⟩
+  Σ (C ≃ D) (λ eq → Is-isomorphism′ a eq x y)  □
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple/Variant.agda 54
-  abstract
-
-    -- Extensionality.
-
-    ext : Extensionality (# 1) (# 1)
-    ext = dependent-extensionality univ₂ (λ _ → univ₁)
-
-    ext₀ : Extensionality (# 0) (# 0)
-    ext₀ = dependent-extensionality univ₁ (λ _ → univ)
+  -- Extensionality.
+
+  ext : Extensionality (# 1) (# 1)
+  ext = dependent-extensionality univ₂ (λ _ → univ₁)
+
+  ext₀ : Extensionality (# 0) (# 0)
+  ext₀ = dependent-extensionality univ₁ (λ _ → univ)
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple/Variant.agda 116
-  abstract
-
-    -- Every element is isomorphic to itself, transported along the
-    -- isomorphism.
-
-    isomorphic-to-itself :
-      (ass : Assumptions) → let open Assumptions ass in
-      ∀ a {B C} (B↔C : B ↔ C) x →
-      Is-isomorphism′ ass a B↔C x (subst (El a) (≃⇒≡ univ (↔⇒≃ B↔C)) x)
-    isomorphic-to-itself ass a B↔C x =
-      transport-theorem (El a) (resp ass a) (resp-id ass a)
-                        univ (↔⇒≃ B↔C) x
-      where open Assumptions ass
-
-    -- Is-isomorphism and Is-isomorphism″ are isomorphic (assuming
-    -- univalence).
-
-    isomorphism-definitions-isomorphic₂ :
-      (ass : Assumptions) →
-      ∀ a {B C} (B↔C : B ↔ C) {x y} →
-      Is-isomorphism a B↔C x y ↔ Is-isomorphism″ ass a B↔C x y
-    isomorphism-definitions-isomorphic₂ ass a B↔C {x} {y} =
-      Is-isomorphism      a B↔C x y  ↝⟨ isomorphism-definitions-isomorphic ass a B↔C ⟩
-      Is-isomorphism′ ass a B↔C x y  ↝⟨ ≡⇒↝ _ $ cong (λ z → z ≡ y) $ isomorphic-to-itself ass a B↔C x ⟩□
-      Is-isomorphism″ ass a B↔C x y  □
+  -- Every element is isomorphic to itself, transported along the
+  -- isomorphism.
+
+  isomorphic-to-itself :
+    (ass : Assumptions) → let open Assumptions ass in
+    ∀ a {B C} (B↔C : B ↔ C) x →
+    Is-isomorphism′ ass a B↔C x (subst (El a) (≃⇒≡ univ (↔⇒≃ B↔C)) x)
+  isomorphic-to-itself ass a B↔C x =
+    transport-theorem (El a) (resp ass a) (resp-id ass a)
+                      univ (↔⇒≃ B↔C) x
+    where open Assumptions ass
+
+  -- Is-isomorphism and Is-isomorphism″ are isomorphic (assuming
+  -- univalence).
+
+  isomorphism-definitions-isomorphic₂ :
+    (ass : Assumptions) →
+    ∀ a {B C} (B↔C : B ↔ C) {x y} →
+    Is-isomorphism a B↔C x y ↔ Is-isomorphism″ ass a B↔C x y
+  isomorphism-definitions-isomorphic₂ ass a B↔C {x} {y} =
+    Is-isomorphism      a B↔C x y  ↝⟨ isomorphism-definitions-isomorphic ass a B↔C ⟩
+    Is-isomorphism′ ass a B↔C x y  ↝⟨ ≡⇒↝ _ $ cong (λ z → z ≡ y) $ isomorphic-to-itself ass a B↔C x ⟩□
+    Is-isomorphism″ ass a B↔C x y  □
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple/Variant.agda 235
-  abstract
-
-    -- The type of isomorphisms between two instances of a structure
-    -- is isomorphic to the type of equalities between the same
-    -- instances (assuming univalence).
-    --
-    -- In short, isomorphism is isomorphic to equality.
-
-    isomorphic↔equal :
-      Assumptions →
-      ∀ c {I₁ I₂} → Isomorphic c I₁ I₂ ↔ (I₁ ≡ I₂)
-    isomorphic↔equal ass c {I₁} {I₂} =
-
-      (∃ λ (C-eq : Carrier c I₁ ↔ Carrier c I₂) →
-         Is-isomorphism (proj₁ c) C-eq (element c I₁) (element c I₂))  ↝⟨ ∃-cong (λ C-eq → isomorphism-definitions-isomorphic₂
-                                                                                             ass (proj₁ c) C-eq) ⟩
-      (∃ λ (C-eq : Carrier c I₁ ↔ Carrier c I₂) →
-         subst (El (proj₁ c)) (≃⇒≡ univ (↔⇒≃ C-eq)) (element c I₁) ≡
-         element c I₂)                                                 ↝⟨ Σ-cong (↔↔≃ ext₀ (proj₂ (proj₁ I₁))) (λ _ → _ □) ⟩
-
-      (∃ λ (C-eq : Carrier c I₁ ≃ Carrier c I₂) →
-         subst (El (proj₁ c)) (≃⇒≡ univ C-eq) (element c I₁) ≡
-         element c I₂)                                                 ↝⟨ inverse $
-                                                                            Σ-cong (≡≃≃ univ) (λ C-eq → ≡⇒↝ _ $ sym $
-                                                                              cong (λ eq → subst (El (proj₁ c)) eq (element c I₁) ≡
-                                                                                           element c I₂)
-                                                                                   (_≃_.left-inverse-of (≡≃≃ univ) C-eq)) ⟩
-      (∃ λ (C-eq : Carrier c I₁ ≡ Carrier c I₂) →
-         subst (El (proj₁ c)) C-eq (element c I₁) ≡ element c I₂)      ↝⟨ inverse $ instances-equal↔ ass c ⟩□
-
-      (I₁ ≡ I₂)                                                        □
-
-      where open Assumptions ass
-
-    -- The first part of the from component of the preceding lemma is
-    -- extensionally equal to a simple function. (The codomain of the
-    -- second part is propositional whenever El (proj₁ c) applied to
-    -- either carrier type is a set.)
-
-    proj₁-from-isomorphic↔equal :
-      ∀ ass c {I J} (I≡J : I ≡ J) →
-      proj₁ (_↔_.from (isomorphic↔equal ass c) I≡J) ≡
-      _≃_.bijection (≡⇒≃ (cong (proj₁ ∘ proj₁) I≡J))
-    proj₁-from-isomorphic↔equal ass (a , P) I≡J =
-      let A = Instance (a , P)
-          B = Σ (∃ (El a)) λ { (C , x) →
-                ∃ λ (S : Is-set C) → proj₁ (P (C , S) x) }
-      in
-
-      cong (_≃_.bijection ∘ ≡⇒≃) (
-        proj₁ (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡ (cong {B = B}
-          (λ { ((C , S) , x , p) → (C , x) , S , p }) I≡J))))            ≡⟨ cong (proj₁ ∘ Σ-≡,≡←≡) $ proj₁-Σ-≡,≡←≡ _ ⟩
-        proj₁ (Σ-≡,≡←≡ (cong proj₁ (cong {B = B}
-          (λ { ((C , S) , x , p) → (C , x) , S , p }) I≡J)))             ≡⟨ cong (proj₁ ∘ Σ-≡,≡←≡) $
-                                                                              cong-∘ {B = B} proj₁
-                                                                                     (λ { ((C , S) , x , p) → (C , x) , S , p }) _ ⟩
-        proj₁ (Σ-≡,≡←≡ (cong {A = A}
-          (λ { ((C , S) , x , p) → C , x }) I≡J))                        ≡⟨ proj₁-Σ-≡,≡←≡ _ ⟩
-        cong proj₁ (cong {A = A} (λ { ((C , S) , x , p) → C , x }) I≡J)  ≡⟨ cong-∘ {A = A} proj₁ (λ { ((C , S) , x , p) → C , x }) _ ⟩∎
-        cong (proj₁ ∘ proj₁) I≡J                                         ∎)
-
-    -- The type of (lifted) isomorphisms between two instances of a
-    -- structure is equal to the type of equalities between the same
-    -- instances (assuming univalence).
-    --
-    -- In short, isomorphism is equal to equality.
-
-    isomorphic≡equal :
-      Assumptions →
-      ∀ c {I₁ I₂} → Isomorphic c I₁ I₂ ≡ (I₁ ≡ I₂)
-    isomorphic≡equal ass c {I₁} {I₂} =
-      ≃⇒≡ univ₁ $ ↔⇒≃ (isomorphic↔equal ass c)
-      where open Assumptions ass
+  -- The type of isomorphisms between two instances of a structure
+  -- is isomorphic to the type of equalities between the same
+  -- instances (assuming univalence).
+  --
+  -- In short, isomorphism is isomorphic to equality.
+
+  isomorphic↔equal :
+    Assumptions →
+    ∀ c {I₁ I₂} → Isomorphic c I₁ I₂ ↔ (I₁ ≡ I₂)
+  isomorphic↔equal ass c {I₁} {I₂} =
+
+    (∃ λ (C-eq : Carrier c I₁ ↔ Carrier c I₂) →
+       Is-isomorphism (proj₁ c) C-eq (element c I₁) (element c I₂))  ↝⟨ ∃-cong (λ C-eq → isomorphism-definitions-isomorphic₂
+                                                                                           ass (proj₁ c) C-eq) ⟩
+    (∃ λ (C-eq : Carrier c I₁ ↔ Carrier c I₂) →
+       subst (El (proj₁ c)) (≃⇒≡ univ (↔⇒≃ C-eq)) (element c I₁) ≡
+       element c I₂)                                                 ↝⟨ Σ-cong (↔↔≃ ext₀ (proj₂ (proj₁ I₁))) (λ _ → _ □) ⟩
+
+    (∃ λ (C-eq : Carrier c I₁ ≃ Carrier c I₂) →
+       subst (El (proj₁ c)) (≃⇒≡ univ C-eq) (element c I₁) ≡
+       element c I₂)                                                 ↝⟨ inverse $
+                                                                          Σ-cong (≡≃≃ univ) (λ C-eq → ≡⇒↝ _ $ sym $
+                                                                            cong (λ eq → subst (El (proj₁ c)) eq (element c I₁) ≡
+                                                                                         element c I₂)
+                                                                                 (_≃_.left-inverse-of (≡≃≃ univ) C-eq)) ⟩
+    (∃ λ (C-eq : Carrier c I₁ ≡ Carrier c I₂) →
+       subst (El (proj₁ c)) C-eq (element c I₁) ≡ element c I₂)      ↝⟨ inverse $ instances-equal↔ ass c ⟩□
+
+    (I₁ ≡ I₂)                                                        □
+
+    where open Assumptions ass
+
+  -- The first part of the from component of the preceding lemma is
+  -- extensionally equal to a simple function. (The codomain of the
+  -- second part is propositional whenever El (proj₁ c) applied to
+  -- either carrier type is a set.)
+
+  proj₁-from-isomorphic↔equal :
+    ∀ ass c {I J} (I≡J : I ≡ J) →
+    proj₁ (_↔_.from (isomorphic↔equal ass c) I≡J) ≡
+    _≃_.bijection (≡⇒≃ (cong (proj₁ ∘ proj₁) I≡J))
+  proj₁-from-isomorphic↔equal ass (a , P) I≡J =
+    let A = Instance (a , P)
+        B = Σ (∃ (El a)) λ { (C , x) →
+              ∃ λ (S : Is-set C) → proj₁ (P (C , S) x) }
+    in
+
+    cong (_≃_.bijection ∘ ≡⇒≃) (
+      proj₁ (Σ-≡,≡←≡ (proj₁ (Σ-≡,≡←≡ (cong {B = B}
+        (λ { ((C , S) , x , p) → (C , x) , S , p }) I≡J))))            ≡⟨ cong (proj₁ ∘ Σ-≡,≡←≡) $ proj₁-Σ-≡,≡←≡ _ ⟩
+      proj₁ (Σ-≡,≡←≡ (cong proj₁ (cong {B = B}
+        (λ { ((C , S) , x , p) → (C , x) , S , p }) I≡J)))             ≡⟨ cong (proj₁ ∘ Σ-≡,≡←≡) $
+                                                                            cong-∘ {B = B} proj₁
+                                                                                   (λ { ((C , S) , x , p) → (C , x) , S , p }) _ ⟩
+      proj₁ (Σ-≡,≡←≡ (cong {A = A}
+        (λ { ((C , S) , x , p) → C , x }) I≡J))                        ≡⟨ proj₁-Σ-≡,≡←≡ _ ⟩
+      cong proj₁ (cong {A = A} (λ { ((C , S) , x , p) → C , x }) I≡J)  ≡⟨ cong-∘ {A = A} proj₁ (λ { ((C , S) , x , p) → C , x }) _ ⟩∎
+      cong (proj₁ ∘ proj₁) I≡J                                         ∎)
+
+  -- The type of (lifted) isomorphisms between two instances of a
+  -- structure is equal to the type of equalities between the same
+  -- instances (assuming univalence).
+  --
+  -- In short, isomorphism is equal to equality.
+
+  isomorphic≡equal :
+    Assumptions →
+    ∀ c {I₁ I₂} → Isomorphic c I₁ I₂ ≡ (I₁ ≡ I₂)
+  isomorphic≡equal ass c {I₁} {I₂} =
+    ≃⇒≡ univ₁ $ ↔⇒≃ (isomorphic↔equal ass c)
+    where open Assumptions ass
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple/Variant.agda 342
-abstract
-
-  -- The cast function respects identities (assuming extensionality).
-
-  cast-id : (ext : Extensionality (# 1) (# 1)) →
-            ∀ a {B} → cast ext a (Eq.id {A = B}) ≡ Eq.id
-  cast-id ext id      = refl _
-  cast-id ext prop    = refl _
-  cast-id ext (k A)   = refl _
-  cast-id ext (a ⇾ b) = lift-equality ext $ cong _≃_.to $
-                          cong₂ (→-cong ext) (cast-id ext a) (cast-id ext b)
-  cast-id ext (a ⊗ b) = lift-equality ext $ cong _≃_.to $
-                          cong₂ _×-cong_ (cast-id ext a) (cast-id ext b)
-  cast-id ext (a ⊕ b) =
-    cast ext a Eq.id ⊎-cong cast ext b Eq.id  ≡⟨ cong₂ _⊎-cong_ (cast-id ext a) (cast-id ext b) ⟩
-    ⟨ [ inj₁ , inj₂ ] , _ ⟩                   ≡⟨ lift-equality ext (ext [ refl ∘ inj₁ , refl ∘ inj₂ ]) ⟩∎
-    Eq.id                                     ∎
+-- The cast function respects identities (assuming extensionality).
+
+cast-id : (ext : Extensionality (# 1) (# 1)) →
+          ∀ a {B} → cast ext a (Eq.id {A = B}) ≡ Eq.id
+cast-id ext id      = refl _
+cast-id ext prop    = refl _
+cast-id ext (k A)   = refl _
+cast-id ext (a ⇾ b) = lift-equality ext $ cong _≃_.to $
+                        cong₂ (→-cong ext) (cast-id ext a) (cast-id ext b)
+cast-id ext (a ⊗ b) = lift-equality ext $ cong _≃_.to $
+                        cong₂ _×-cong_ (cast-id ext a) (cast-id ext b)
+cast-id ext (a ⊕ b) =
+  cast ext a Eq.id ⊎-cong cast ext b Eq.id  ≡⟨ cong₂ _⊎-cong_ (cast-id ext a) (cast-id ext b) ⟩
+  ⟨ [ inj₁ , inj₂ ] , _ ⟩                   ≡⟨ lift-equality ext (ext [ refl ∘ inj₁ , refl ∘ inj₂ ]) ⟩∎
+  Eq.id                                     ∎
hunk ./Univalence-axiom/Isomorphism-is-equality/Simple/Variant.agda 378
-abstract
-
-  -- The two definitions of "being an isomorphism" are "isomorphic"
-  -- (in bijective correspondence), assuming univalence.
-
-  isomorphism-definitions-isomorphic :
-    (ass : Assumptions) → let open Assumptions ass in
-    ∀ a {B C} (B↔C : B ↔ C) {x y} →
-    Is-isomorphism a B↔C x y ↔ Is-isomorphism′ ext a B↔C x y
-
-  isomorphism-definitions-isomorphic ass id B↔C {x} {y} =
-
-    (_↔_.to B↔C x ≡ y)  □
-
-  isomorphism-definitions-isomorphic ass prop B↔C {P} {Q} =
-
-    ↑ _ (proj₁ P ⇔ proj₁ Q)  ↝⟨ ↑↔ ⟩
-
-    (proj₁ P ⇔ proj₁ Q)      ↝⟨ ⇔↔≃ ext₀ (proj₂ P) (proj₂ Q) ⟩
-
-    (proj₁ P ≃ proj₁ Q)      ↔⟨ inverse $ ≡≃≃ univ ⟩
-
-    (proj₁ P ≡ proj₁ Q)      ↝⟨ ignore-propositional-component (H-level-propositional ext₀ 1) ⟩□
-
-    (P ≡ Q)                □
-
-    where open Assumptions ass
-
-  isomorphism-definitions-isomorphic ass (k A) B↔C {x} {y} =
-
-    (x ≡ y) □
-
-  isomorphism-definitions-isomorphic ass (a ⇾ b) B↔C {f} {g} =
-    let B≃C = ↔⇒≃ B↔C in
-
-    (∀ x y → Is-isomorphism a B↔C x y →
-             Is-isomorphism b B↔C (f x) (g y))                     ↔⟨ ∀-preserves ext (λ _ → ∀-preserves ext λ _ → ↔⇒≃ $
-                                                                        →-cong ext (isomorphism-definitions-isomorphic ass a B↔C)
-                                                                                   (isomorphism-definitions-isomorphic ass b B↔C)) ⟩
-    (∀ x y → to (cast ext a B≃C) x ≡ y →
-             to (cast ext b B≃C) (f x) ≡ g y)                      ↔⟨ inverse $ ∀-preserves ext (λ x →
-                                                                        ↔⇒≃ $ ∀-intro ext (λ y _ → to (cast ext b B≃C) (f x) ≡ g y)) ⟩
-    (∀ x → to (cast ext b B≃C) (f x) ≡ g (to (cast ext a B≃C) x))  ↔⟨ extensionality-isomorphism ext ⟩
-
-    (to (cast ext b B≃C) ∘ f ≡ g ∘ to (cast ext a B≃C))            ↝⟨ inverse $ ∘from≡↔≡∘to ext (cast ext a B≃C) ⟩□
-
-    (to (cast ext b B≃C) ∘ f ∘ from (cast ext a B≃C) ≡ g)          □
-
-    where
-    open _≃_
-    open Assumptions ass
-
-  isomorphism-definitions-isomorphic ass (a ⊕ b) B↔C {inj₁ x} {inj₁ y} =
-    let B≃C = ↔⇒≃ B↔C in
-
-    Is-isomorphism a B↔C x y                 ↝⟨ isomorphism-definitions-isomorphic ass a B↔C ⟩
-
-    (to (cast ext a B≃C) x ≡ y)              ↝⟨ ≡↔inj₁≡inj₁ ⟩□
-
-    (inj₁ (to (cast ext a B≃C) x) ≡ inj₁ y)  □
-
-    where
-    open _≃_
-    open Assumptions ass
-
-  isomorphism-definitions-isomorphic ass (a ⊕ b) B↔C {inj₂ x} {inj₂ y} =
-    let B≃C = ↔⇒≃ B↔C in
-
-    Is-isomorphism b B↔C x y                 ↝⟨ isomorphism-definitions-isomorphic ass b B↔C ⟩
-
-    (to (cast ext b B≃C) x ≡ y)              ↝⟨ ≡↔inj₂≡inj₂ ⟩□
-
-    (inj₂ (to (cast ext b B≃C) x) ≡ inj₂ y)  □
-
-    where
-    open _≃_
-    open Assumptions ass
-
-  isomorphism-definitions-isomorphic ass (a ⊕ b) B↔C {inj₁ x} {inj₂ y} =
-
-    ⊥                  ↝⟨ ⊥↔uninhabited ⊎.inj₁≢inj₂ ⟩□
-
-    (inj₁ _ ≡ inj₂ _)  □
-
-  isomorphism-definitions-isomorphic ass (a ⊕ b) B↔C {inj₂ x} {inj₁ y} =
-
-    ⊥                  ↝⟨ ⊥↔uninhabited (⊎.inj₁≢inj₂ ∘ sym) ⟩□
-
-    (inj₂ _ ≡ inj₁ _)  □
-
-  isomorphism-definitions-isomorphic ass (a ⊗ b) B↔C {x , u} {y , v} =
-    let B≃C = ↔⇒≃ B↔C in
-
-    Is-isomorphism  a B↔C x y × Is-isomorphism  b B↔C u v        ↝⟨ isomorphism-definitions-isomorphic ass a B↔C ×-cong
-                                                                    isomorphism-definitions-isomorphic ass b B↔C ⟩
-    (to (cast ext a B≃C) x ≡ y × to (cast ext b B≃C) u ≡ v)      ↝⟨ ≡×≡↔≡ ⟩□
-
-    ((to (cast ext a B≃C) x , to (cast ext b B≃C) u) ≡ (y , v))  □
-
-    where
-    open _≃_
-    open Assumptions ass
+-- The two definitions of "being an isomorphism" are "isomorphic"
+-- (in bijective correspondence), assuming univalence.
+
+isomorphism-definitions-isomorphic :
+  (ass : Assumptions) → let open Assumptions ass in
+  ∀ a {B C} (B↔C : B ↔ C) {x y} →
+  Is-isomorphism a B↔C x y ↔ Is-isomorphism′ ext a B↔C x y
+
+isomorphism-definitions-isomorphic ass id B↔C {x} {y} =
+
+  (_↔_.to B↔C x ≡ y)  □
+
+isomorphism-definitions-isomorphic ass prop B↔C {P} {Q} =
+
+  ↑ _ (proj₁ P ⇔ proj₁ Q)  ↝⟨ ↑↔ ⟩
+
+  (proj₁ P ⇔ proj₁ Q)      ↝⟨ ⇔↔≃ ext₀ (proj₂ P) (proj₂ Q) ⟩
+
+  (proj₁ P ≃ proj₁ Q)      ↔⟨ inverse $ ≡≃≃ univ ⟩
+
+  (proj₁ P ≡ proj₁ Q)      ↝⟨ ignore-propositional-component (H-level-propositional ext₀ 1) ⟩□
+
+  (P ≡ Q)                □
+
+  where open Assumptions ass
+
+isomorphism-definitions-isomorphic ass (k A) B↔C {x} {y} =
+
+  (x ≡ y) □
+
+isomorphism-definitions-isomorphic ass (a ⇾ b) B↔C {f} {g} =
+  let B≃C = ↔⇒≃ B↔C in
+
+  (∀ x y → Is-isomorphism a B↔C x y →
+           Is-isomorphism b B↔C (f x) (g y))                     ↔⟨ ∀-preserves ext (λ _ → ∀-preserves ext λ _ → ↔⇒≃ $
+                                                                      →-cong ext (isomorphism-definitions-isomorphic ass a B↔C)
+                                                                                 (isomorphism-definitions-isomorphic ass b B↔C)) ⟩
+  (∀ x y → to (cast ext a B≃C) x ≡ y →
+           to (cast ext b B≃C) (f x) ≡ g y)                      ↔⟨ inverse $ ∀-preserves ext (λ x →
+                                                                      ↔⇒≃ $ ∀-intro ext (λ y _ → to (cast ext b B≃C) (f x) ≡ g y)) ⟩
+  (∀ x → to (cast ext b B≃C) (f x) ≡ g (to (cast ext a B≃C) x))  ↔⟨ extensionality-isomorphism ext ⟩
+
+  (to (cast ext b B≃C) ∘ f ≡ g ∘ to (cast ext a B≃C))            ↝⟨ inverse $ ∘from≡↔≡∘to ext (cast ext a B≃C) ⟩□
+
+  (to (cast ext b B≃C) ∘ f ∘ from (cast ext a B≃C) ≡ g)          □
+
+  where
+  open _≃_
+  open Assumptions ass
+
+isomorphism-definitions-isomorphic ass (a ⊕ b) B↔C {inj₁ x} {inj₁ y} =
+  let B≃C = ↔⇒≃ B↔C in
+
+  Is-isomorphism a B↔C x y                 ↝⟨ isomorphism-definitions-isomorphic ass a B↔C ⟩
+
+  (to (cast ext a B≃C) x ≡ y)              ↝⟨ ≡↔inj₁≡inj₁ ⟩□
+
+  (inj₁ (to (cast ext a B≃C) x) ≡ inj₁ y)  □
+
+  where
+  open _≃_
+  open Assumptions ass
+
+isomorphism-definitions-isomorphic ass (a ⊕ b) B↔C {inj₂ x} {inj₂ y} =
+  let B≃C = ↔⇒≃ B↔C in
+
+  Is-isomorphism b B↔C x y                 ↝⟨ isomorphism-definitions-isomorphic ass b B↔C ⟩
+
+  (to (cast ext b B≃C) x ≡ y)              ↝⟨ ≡↔inj₂≡inj₂ ⟩□
+
+  (inj₂ (to (cast ext b B≃C) x) ≡ inj₂ y)  □
+
+  where
+  open _≃_
+  open Assumptions ass
+
+isomorphism-definitions-isomorphic ass (a ⊕ b) B↔C {inj₁ x} {inj₂ y} =
+
+  ⊥                  ↝⟨ ⊥↔uninhabited ⊎.inj₁≢inj₂ ⟩□
+
+  (inj₁ _ ≡ inj₂ _)  □
+
+isomorphism-definitions-isomorphic ass (a ⊕ b) B↔C {inj₂ x} {inj₁ y} =
+
+  ⊥                  ↝⟨ ⊥↔uninhabited (⊎.inj₁≢inj₂ ∘ sym) ⟩□
+
+  (inj₂ _ ≡ inj₁ _)  □
+
+isomorphism-definitions-isomorphic ass (a ⊗ b) B↔C {x , u} {y , v} =
+  let B≃C = ↔⇒≃ B↔C in
+
+  Is-isomorphism  a B↔C x y × Is-isomorphism  b B↔C u v        ↝⟨ isomorphism-definitions-isomorphic ass a B↔C ×-cong
+                                                                  isomorphism-definitions-isomorphic ass b B↔C ⟩
+  (to (cast ext a B≃C) x ≡ y × to (cast ext b B≃C) u ≡ v)      ↝⟨ ≡×≡↔≡ ⟩□
+
+  ((to (cast ext a B≃C) x , to (cast ext b B≃C) u) ≡ (y , v))  □
+
+  where
+  open _≃_
+  open Assumptions ass

Context:

[Added dependent-cong.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150612104142
 Ignore-this: 1c7d1a2700b43e9a62f78f25e01aca3a
] 
[Added Σ-closure-contractible.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150529195443
 Ignore-this: 3dfdb65e5607808d5dacc8ae504e9980
] 
[Added swap-as-an-equality and swap≢refl.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150522215632
 Ignore-this: 644722ecd5aa2f11a26f88c85056e737
] 
[Added [subst≡]≡[trans≡trans].
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150522211736
 Ignore-this: d779a471c9c3f9c6e9a3ccc87691223
] 
[Added ↑⁻¹-closure.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150521210811
 Ignore-this: 4a98580fa599f0621a1b7a5d86645e51
] 
[Added more lemmas related to h-levels for Is-equivalence and _≃_.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150515184045
 Ignore-this: 5359663b4a86fbc3a292079e0078116f
] 
[Added subst-as-equivalence, redefined subst-is-equivalence.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150512172159
 Ignore-this: 405312a6394ae9b1faa4651c17154be1
] 
[Added ≃-to-≡↔≡ and ≃-from-≡↔≡.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150508180753
 Ignore-this: 7c578c35e306e4c399f24de24b0497d5
] 
[Added inverse-involutive and inverse-isomorphism.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150508180721
 Ignore-this: cd33d440f1dcb4c35681593b71d3b740
] 
[Added subst-good-ext.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150423171035
 Ignore-this: c8f8378a406ed99498cf2ff1837b58e1
] 
[Added not≡⇒≢.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150423122055
 Ignore-this: 59a567149b35b942ac1b256b9d71c941
] 
[Made some abstract definitions concrete.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150422174856
 Ignore-this: 7ad1676b320e8671553c85db3df4f7ee
] 
[Added with-lower-level and prop-elim.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150422153652
 Ignore-this: d7b3b039ba73cad7347e0687c10940d0
] 
[Added constant-endofunction⇔h-stable along with supporting lemmas.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150417205500
 Ignore-this: 463a0af5dcc9f0f011992eb1faace42b
] 
[Added H-level.Truncation.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150417194613
 Ignore-this: 5ebfded1c68f27626ee7999161fa499d
] 
[Fixed an outdated comment.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150417185858
 Ignore-this: be0a6be2ca0e18486e268e916586df2e
] 
[Renamed idempotent to involutive.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150416180529
 Ignore-this: b932f1aca37e82c7394b94bc005d6324
] 
[Added ≡-preserves-≃.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150415094528
 Ignore-this: dfa8624442d9a77a94832c58d82f4f89
] 
[Added proj₂-closure.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150409164405
 Ignore-this: 1eb8286833bc2ae912efffc84c44c618
] 
[Added ≡⇒→-≃⇒≡.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150402174807
 Ignore-this: 19d4450c7f704afc00301360266b9be1
] 
[Added to-implication-≡⇒↝.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150402174517
 Ignore-this: 492a5b6fc05ba3ccba6e63e8d1d5d2b6
] 
[Added Bool.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150331151629
 Ignore-this: f4d8eff18e1d71315cc6443ae84f6b78
] 
[Removed contractible↔⊤, made contractible⇔⊤↔ non-abstract.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150327105006
 Ignore-this: 69fe878f8b3cccdc7575e321d92c9637
] 
[Added the pattern synonyms yes and no.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150326193510
 Ignore-this: ea59bc517ac8718ef38f325014aa7150
] 
[Replaced subst-in-terms-of-≡⇒≃ with subst-in-terms-of-≡⇒↝.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150326192608
 Ignore-this: f4142235aea24832ddeb05f207b397e2
 * Also replaced subst-in-terms-of-from∘≡⇒≃ with
   subst-in-terms-of-inverse∘≡⇒↝.
] 
[Added the pattern synonyms fzero and fsuc.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150326192125
 Ignore-this: cc53c728701046d61fbf0c08c733a032
] 
[Added proj₁-closure.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150326111900
 Ignore-this: 3b4914e42880128f8cfc3ec121f2250a
] 
[Added subst-application.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150325173324
 Ignore-this: efaacdb40a30f0074093074afe3bfeb4
] 
[Added lots of parentheses and a fixity declaration.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20150325164616
 Ignore-this: ec080caa7d511c6830182bb24fe08504
 * Agda's default fixity has been changed to "unrelated".
] 
[Made the code type-check using the current development version of Agda.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20141128233904
 Ignore-this: 1cbc7b7c79495d8dcddeb34d581e5b1b
] 
[Added propositional-domain⇒constant and propositional-identity⇒set.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20140404184157
 Ignore-this: e3353032d249195bbf28269321cf27f4
] 
[Moved decidable⇒set from H-level to Equality.Decidable-UIP.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20140404184058
 Ignore-this: 58eb401083246a18eeb0b917b5130dda
] 
[Renamed constant to decidable⇒constant.
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20140404174204
 Ignore-this: 70cd30b2649b0f76e723723467338140
] 
[TAG Code corresponding to "Isomorphism is equality" (2013-10-01)
Nils Anders Danielsson <nils.anders.danielsson at gmail.com>**20131001155219
 Ignore-this: 27b7ac73d249cacf2206caabfba4f4bf
] 
Patch bundle hash:
82befe52a10c13b6cfcb487ddb74408c9aa1ac1b


More information about the Agda-dev mailing list