88
99module Axiom.UniquenessOfIdentityProofs where
1010
11- open import Data.Bool.Base using (true; false)
12- open import Data.Empty
13- open import Relation.Nullary.Reflects using (invert)
14- open import Relation.Nullary hiding (Irrelevant)
11+ open import Level using (Level)
12+ open import Relation.Nullary.Decidable.Core using (recompute; recompute-irr)
1513open import Relation.Binary.Core
1614open import Relation.Binary.Definitions
1715open import Relation.Binary.PropositionalEquality.Core
1816open import Relation.Binary.PropositionalEquality.Properties
1917
18+ private
19+ variable
20+ a : Level
21+ A : Set a
22+ x y : A
23+
2024------------------------------------------------------------------------
2125-- Definition
2226--
2327-- Uniqueness of Identity Proofs (UIP) states that all proofs of
2428-- equality are themselves equal. In other words, the equality relation
2529-- is irrelevant. Here we define UIP relative to a given type.
2630
27- UIP : ∀ {a} (A : Set a) → Set a
31+ UIP : (A : Set a) → Set a
2832UIP A = Irrelevant {A = A} _≡_
2933
3034------------------------------------------------------------------------
@@ -39,11 +43,11 @@ UIP A = Irrelevant {A = A} _≡_
3943-- the image of any other proof.
4044
4145module Constant⇒UIP
42- {a} {A : Set a} (f : _≡_ {A = A} ⇒ _≡_)
43- (f-constant : ∀ {a b } (p q : a ≡ b ) → f p ≡ f q)
46+ (f : _≡_ {A = A} ⇒ _≡_)
47+ (f-constant : ∀ {x y } (p q : x ≡ y ) → f p ≡ f q)
4448 where
4549
46- ≡-canonical : ∀ {a b} (p : a ≡ b ) → trans (sym (f refl)) (f p) ≡ p
50+ ≡-canonical : (p : x ≡ y ) → trans (sym (f refl)) (f p) ≡ p
4751 ≡-canonical refl = trans-symˡ (f refl)
4852
4953 ≡-irrelevant : UIP A
@@ -59,19 +63,13 @@ module Constant⇒UIP
5963-- function over proofs of equality which is constant: it returns the
6064-- proof produced by the decision procedure.
6165
62- module Decidable⇒UIP
63- {a} {A : Set a} (_≟_ : Decidable {A = A} _≡_)
64- where
66+ module Decidable⇒UIP (_≟_ : Decidable {A = A} _≡_) where
6567
6668 ≡-normalise : _≡_ {A = A} ⇒ _≡_
67- ≡-normalise {a} {b} a≡b with a ≟ b
68- ... | true because [p] = invert [p]
69- ... | false because [¬p] = ⊥-elim (invert [¬p] a≡b)
70-
71- ≡-normalise-constant : ∀ {a b} (p q : a ≡ b) → ≡-normalise p ≡ ≡-normalise q
72- ≡-normalise-constant {a} {b} p q with a ≟ b
73- ... | true because _ = refl
74- ... | false because [¬p] = ⊥-elim (invert [¬p] p)
69+ ≡-normalise {x} {y} x≡y = recompute (x ≟ y) x≡y
70+
71+ ≡-normalise-constant : (p q : x ≡ y) → ≡-normalise p ≡ ≡-normalise q
72+ ≡-normalise-constant {x = x} {y = y} = recompute-irr (x ≟ y)
7573
7674 ≡-irrelevant : UIP A
7775 ≡-irrelevant = Constant⇒UIP.≡-irrelevant ≡-normalise ≡-normalise-constant
0 commit comments