{-# OPTIONS --without-K #-}
open import Equality
module Equality.Decidable-UIP
{reflexive} (eq : ∀ {a p} → Equality-with-J a p reflexive) where
open Derived-definitions-and-properties eq
open import Logical-equivalence using (module _⇔_)
open import H-level eq
open import Prelude
Constant : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → Set (a ⊔ b)
Constant f = ∀ x y → f x ≡ f y
_Left-inverse-of_ : ∀ {a b} {A : Set a} {B : Set b} →
(B → A) → (A → B) → Set a
g Left-inverse-of f = ∀ x → g (f x) ≡ x
abstract
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 ∎
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ʳ _)
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))
decidable⇒constant : ∀ {a} {A : Set a} → Dec A →
∃ λ (f : A → A) → Constant f
decidable⇒constant (inj₁ x) = (const x , λ _ _ → refl x)
decidable⇒constant (inj₂ ¬x) = (id , λ _ → ⊥-elim ∘ ¬x)
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))
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)
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)
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) _ _)