forked from agda/agda-stdlib
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathReflects.agda
100 lines (77 loc) · 3.54 KB
/
Reflects.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
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
------------------------------------------------------------------------
-- The Agda standard library
--
-- Properties of the `Reflects` construct
------------------------------------------------------------------------
{-# OPTIONS --cubical-compatible --safe #-}
module Relation.Nullary.Reflects where
open import Agda.Builtin.Equality
open import Data.Bool.Base
open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
open import Data.Product.Base using (_×_; _,_; proj₁; proj₂)
open import Level using (Level)
open import Function.Base using (id; _$_; _∘_; const)
open import Relation.Nullary.Negation.Core using (¬_; contradiction; _¬-⊎_)
private
variable
p q : Level
P Q : Set p
------------------------------------------------------------------------
-- `Reflects` idiom.
-- The truth value of P is reflected by a boolean value.
-- `Reflects P b` is equivalent to `if b then P else ¬ P`.
data Reflects {p} (P : Set p) : Bool → Set p where
ofʸ : ( p : P) → Reflects P true
ofⁿ : (¬p : ¬ P) → Reflects P false
------------------------------------------------------------------------
-- Constructors and destructors
T⁺ : ∀ b → Reflects (T b) b
T⁺ true = ofʸ _
T⁺ false = ofⁿ id
-- These lemmas are intended to be used mostly when `b` is a value, so
-- that the `if` expressions have already been evaluated away.
-- In this case, `of` works like the relevant constructor (`ofⁿ` or
-- `ofʸ`), and `invert` strips off the constructor to just give either
-- the proof of `P` or the proof of `¬ P`.
of : ∀ {b} → if b then P else ¬ P → Reflects P b
of {b = false} ¬p = ofⁿ ¬p
of {b = true } p = ofʸ p
invert : ∀ {b} → Reflects P b → if b then P else ¬ P
invert (ofʸ p) = p
invert (ofⁿ ¬p) = ¬p
------------------------------------------------------------------------
-- Interaction with negation, product, sums etc.
-- If we can decide P, then we can decide its negation.
¬-reflects : ∀ {b} → Reflects P b → Reflects (¬ P) (not b)
¬-reflects (ofʸ p) = ofⁿ (_$ p)
¬-reflects (ofⁿ ¬p) = ofʸ ¬p
-- If we can decide P and Q then we can decide their product
infixr 2 _×-reflects_
_×-reflects_ : ∀ {a b} → Reflects P a → Reflects Q b →
Reflects (P × Q) (a ∧ b)
ofʸ p ×-reflects ofʸ q = ofʸ (p , q)
ofʸ p ×-reflects ofⁿ ¬q = ofⁿ (¬q ∘ proj₂)
ofⁿ ¬p ×-reflects _ = ofⁿ (¬p ∘ proj₁)
infixr 1 _⊎-reflects_
_⊎-reflects_ : ∀ {a b} → Reflects P a → Reflects Q b →
Reflects (P ⊎ Q) (a ∨ b)
ofʸ p ⊎-reflects _ = ofʸ (inj₁ p)
ofⁿ ¬p ⊎-reflects ofʸ q = ofʸ (inj₂ q)
ofⁿ ¬p ⊎-reflects ofⁿ ¬q = ofⁿ (¬p ¬-⊎ ¬q)
infixr 2 _→-reflects_
_→-reflects_ : ∀ {a b} → Reflects P a → Reflects Q b →
Reflects (P → Q) (not a ∨ b)
ofʸ p →-reflects ofʸ q = ofʸ (const q)
ofʸ p →-reflects ofⁿ ¬q = ofⁿ (¬q ∘ (_$ p))
ofⁿ ¬p →-reflects _ = ofʸ (λ p → contradiction p ¬p)
------------------------------------------------------------------------
-- Other lemmas
fromEquivalence : ∀ {b} → (T b → P) → (P → T b) → Reflects P b
fromEquivalence {b = true} sound complete = ofʸ (sound _)
fromEquivalence {b = false} sound complete = ofⁿ complete
-- `Reflects` is deterministic.
det : ∀ {b b′} → Reflects P b → Reflects P b′ → b ≡ b′
det (ofʸ p) (ofʸ p′) = refl
det (ofʸ p) (ofⁿ ¬p′) = contradiction p ¬p′
det (ofⁿ ¬p) (ofʸ p′) = contradiction p′ ¬p
det (ofⁿ ¬p) (ofⁿ ¬p′) = refl