module 1Lab.Counterexamples.IsIso where
is-iso is not a proposition🔗
We show that if is-iso
were a proposition, then
(x : A) → x ≡ x
would be contractible for any choice of
A
. Taking A
to be S¹
, we show that this can not be the
case. Suppose that is-iso is a proposition.
module
_ (iso-is-prop : ∀ {ℓ} {A B : Type ℓ} {f : A → B} → is-prop (is-iso f))
where
First we characterise the type is-iso f
by showing that, if it is inhabited,
then it is equivalent to the centre of A
, i.e. the
loop-assigning maps of A
:
: ∀ {ℓ} {A B : Type ℓ} {f : A → B} → is-iso f → is-iso f ≃ ((x : A) → x ≡ x)
lemma {A = A} {B} {f} iiso =
lemma (λ _ f → is-iso (f .fst) ≃ ((x : A) → x ≡ x))
EquivJ (Iso→Equiv helper)
(f , is-iso→is-equiv iiso)
where
By equivalence
induction, it suffices to cover the case where f
is the
identity function. In that case, we can construct an isomorphism quite
readily, where the proof uses our assumption iso-is-prop
for convenience.
: Iso _ _
helper .fst iiso x =
helper (iiso .is-iso.linv x) ∙ iiso .is-iso.rinv x
sym .snd .is-iso.inv x = iso (λ x → x) x (λ _ → refl)
helper .snd .is-iso.rinv p = funext λ x → ∙-idl _
helper .snd .is-iso.linv x = iso-is-prop _ _ helper
We thus have that is-iso id ≃ ((x : A) → x ≡ x)
- since
the former is a prop (by assumption), then so is the latter:
: ∀ {ℓ} {A : Type ℓ} → is-prop ((x : A) → x ≡ x)
is-prop-loops {A = A} = equiv→is-hlevel 1 (helper .fst) (helper .snd) iso-is-prop
is-prop-loops where helper = lemma {f = λ (x : A) → x}
(iso (λ x → x) (λ x → refl) (λ x → refl))
Thus, it suffices to choose a type for which
(x : A) → x ≡ x
has two distinct elements. We go with the
circle, S¹
:
: ¬ is-prop ((x : S¹) → x ≡ x)
¬is-prop-loops = refl≠loop $
¬is-prop-loops prop (prop (λ x → refl)
happly λ { base → loop
; (loop i) j → hcomp (∂ i ∨ ∂ j) λ where
(k = i0) → loop (i ∨ j)
k (i = i0) → loop j
k (i = i1) → loop (k ∧ j)
k (j = i0) → loop i
k (j = i1) → loop (k ∧ i)
k })
base
Hence, a contradiction:
: ⊥
contra = ¬is-prop-loops is-prop-loops contra