forked from agda/agda-stdlib
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathUniquenessOfIdentityProofs.agda
76 lines (59 loc) · 2.64 KB
/
UniquenessOfIdentityProofs.agda
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
------------------------------------------------------------------------
-- The Agda standard library
--
-- Results concerning uniqueness of identity proofs
------------------------------------------------------------------------
{-# OPTIONS --cubical-compatible --safe #-}
module Axiom.UniquenessOfIdentityProofs where
open import Level using (Level)
open import Relation.Nullary.Decidable.Core using (recompute; recompute-constant)
open import Relation.Binary.Core
open import Relation.Binary.Definitions
open import Relation.Binary.PropositionalEquality.Core
open import Relation.Binary.PropositionalEquality.Properties
private
variable
a : Level
A : Set a
x y : A
------------------------------------------------------------------------
-- Definition
--
-- Uniqueness of Identity Proofs (UIP) states that all proofs of
-- equality are themselves equal. In other words, the equality relation
-- is irrelevant. Here we define UIP relative to a given type.
UIP : (A : Set a) → Set a
UIP A = Irrelevant {A = A} _≡_
------------------------------------------------------------------------
-- Properties
-- UIP always holds when using axiom K
-- (see `Axiom.UniquenessOfIdentityProofs.WithK`).
-- The existence of a constant function over proofs of equality for
-- elements in A is enough to prove UIP for A. Indeed, we can relate any
-- proof to its image via this function which we then know is equal to
-- the image of any other proof.
module Constant⇒UIP
(f : _≡_ {A = A} ⇒ _≡_)
(f-constant : ∀ {x y} (p q : x ≡ y) → f p ≡ f q)
where
≡-canonical : (p : x ≡ y) → trans (sym (f refl)) (f p) ≡ p
≡-canonical refl = trans-symˡ (f refl)
≡-irrelevant : UIP A
≡-irrelevant p q = begin
p ≡⟨ sym (≡-canonical p) ⟩
trans (sym (f refl)) (f p) ≡⟨ cong (trans _) (f-constant p q) ⟩
trans (sym (f refl)) (f q) ≡⟨ ≡-canonical q ⟩
q ∎
where open ≡-Reasoning
-- If equality is decidable for a given type, then we can prove UIP for
-- that type. Indeed, the decision procedure allows us to define a
-- function over proofs of equality which is constant: it returns the
-- proof produced by the decision procedure.
module Decidable⇒UIP (_≟_ : DecidableEquality A)
where
≡-normalise : _≡_ {A = A} ⇒ _≡_
≡-normalise {x} {y} x≡y = recompute (x ≟ y) x≡y
≡-normalise-constant : (p q : x ≡ y) → ≡-normalise p ≡ ≡-normalise q
≡-normalise-constant {x = x} {y = y} = recompute-constant (x ≟ y)
≡-irrelevant : UIP A
≡-irrelevant = Constant⇒UIP.≡-irrelevant ≡-normalise ≡-normalise-constant