Skip to content

Commit b13a032

Browse files
authored
Adds Relation.Nullary.Recomputable plus consequences (#2243)
* prototype for fixing #2199 * delegate to `Relation.Nullary.Negation.Core.weak-contradiction` * refactoring: lift out `Recomputable` in its own right * forgot to add this module * refactoring: tweak * tidying up; added `CHANGELOG` * more tidying up * streamlined `import`s * removed `Recomputable` from `Relation.Nullary` * fixed multiple definitions in `Relation.Unary` * reorder `CHANGELOG` entries after merge * `contradiciton` via `weak-contradiction` * irrefutable `with` * use `of` * only use *irrelevant* `⊥-elim-irr` * tightened `import`s * removed `irr-contradiction` * inspired by #2329 * conjecture: all uses of `contradiction` can be made weak * second thoughts: reverted last round of chnages * lazier pattern analysis cf. #2055 * dependencies: uncoupled from `Recomputable` * moved `⊥` and `¬ A` properties to this one place * removed `contradictionᵒ` and rephrased everything in terms of `weak-contradiction` * knock-on consequences; simplified `import`s * narrow `import`s * narrow `import`s; irrefutable `with` * narrow `import`s; `CHANGELOG` * narrow `import`s * response to review comments * irrelevance of `recompute` * knock-on, plus proof of `UIP` from #2149 * knock-ons from renaming * knock-on from renaming * pushed proof `recompute-constant` to `Recomputable`
1 parent d970b78 commit b13a032

12 files changed

+207
-107
lines changed

CHANGELOG.md

+53-25
Original file line numberDiff line numberDiff line change
@@ -118,25 +118,9 @@ New modules
118118
Algebra.Module.Bundles.Raw
119119
```
120120

121-
* Properties of `List` modulo `Setoid` equality (currently only the ([],++) monoid):
122-
```
123-
Data.List.Relation.Binary.Equality.Setoid.Properties
124-
```
125-
126-
* Refactoring of `Data.List.Base.{scanr|scanl}` and their properties:
127-
```
128-
Data.List.Scans.Base
129-
Data.List.Scans.Properties
130-
```
131-
132-
* Prime factorisation of natural numbers.
133-
```
134-
Data.Nat.Primality.Factorisation
135-
```
136-
137-
* Consequences of 'infinite descent' for (accessible elements of) well-founded relations:
121+
* Nagata's construction of the "idealization of a module":
138122
```agda
139-
Induction.InfiniteDescent
123+
Algebra.Module.Construct.Idealization
140124
```
141125

142126
* The unique morphism from the initial, resp. terminal, algebra:
@@ -145,16 +129,35 @@ New modules
145129
Algebra.Morphism.Construct.Terminal
146130
```
147131

148-
* Nagata's construction of the "idealization of a module":
132+
* Pointwise and equality relations over indexed containers:
149133
```agda
150-
Algebra.Module.Construct.Idealization
134+
Data.Container.Indexed.Relation.Binary.Pointwise
135+
Data.Container.Indexed.Relation.Binary.Pointwise.Properties
136+
Data.Container.Indexed.Relation.Binary.Equality.Setoid
137+
```
138+
139+
* Refactoring of `Data.List.Base.{scanr|scanl}` and their properties:
140+
```
141+
Data.List.Scans.Base
142+
Data.List.Scans.Properties
143+
```
144+
145+
* Properties of `List` modulo `Setoid` equality (currently only the ([],++) monoid):
146+
```
147+
Data.List.Relation.Binary.Equality.Setoid.Properties
151148
```
152149

153150
* `Data.List.Relation.Binary.Sublist.Propositional.Slice`
154151
replacing `Data.List.Relation.Binary.Sublist.Propositional.Disjoint` (*)
155152
and adding new functions:
156153
- `⊆-upper-bound-is-cospan` generalising `⊆-disjoint-union-is-cospan` from (*)
157154
- `⊆-upper-bound-cospan` generalising `⊆-disjoint-union-cospan` from (*)
155+
```
156+
157+
* Prime factorisation of natural numbers.
158+
```
159+
Data.Nat.Primality.Factorisation
160+
```
158161
159162
* `Data.Vec.Functional.Relation.Binary.Permutation`, defining:
160163
```agda
@@ -180,6 +183,11 @@ New modules
180183
_⇨_ = setoid
181184
```
182185

186+
* Consequences of 'infinite descent' for (accessible elements of) well-founded relations:
187+
```agda
188+
Induction.InfiniteDescent
189+
```
190+
183191
* Symmetric interior of a binary relation
184192
```
185193
Relation.Binary.Construct.Interior.Symmetric
@@ -190,12 +198,11 @@ New modules
190198
Relation.Binary.Properties.DecSetoid
191199
```
192200

193-
* Pointwise and equality relations over indexed containers:
201+
* Systematise the use of `Recomputable A = .A → A`:
194202
```agda
195-
Data.Container.Indexed.Relation.Binary.Pointwise
196-
Data.Container.Indexed.Relation.Binary.Pointwise.Properties
197-
Data.Container.Indexed.Relation.Binary.Equality.Setoid
203+
Relation.Nullary.Recomputable
198204
```
205+
with `Recomputable` exported publicly from `Relation.Nullary`.
199206

200207
* New IO primitives to handle buffering
201208
```agda
@@ -372,6 +379,11 @@ Additions to existing modules
372379
Subtrees o c = (r : Response c) → X (next c r)
373380
```
374381

382+
* In `Data.Empty`:
383+
```agda
384+
⊥-elim-irr : .⊥ → Whatever
385+
```
386+
375387
* In `Data.Fin.Properties`:
376388
```agda
377389
nonZeroIndex : Fin n → ℕ.NonZero n
@@ -674,7 +686,7 @@ Additions to existing modules
674686
* Added new definitions in `Relation.Binary.Definitions`
675687
```
676688
Stable _∼_ = ∀ x y → Nullary.Stable (x ∼ y)
677-
Empty _∼_ = ∀ {x y} → x ∼ y → ⊥
689+
Empty _∼_ = ∀ {x y} → ¬ (x ∼ y)
678690
```
679691

680692
* Added new proofs in `Relation.Binary.Properties.Setoid`:
@@ -690,11 +702,27 @@ Additions to existing modules
690702
WeaklyDecidable : Set _
691703
```
692704

705+
* Added new proof in `Relation.Nullary.Decidable.Core`:
706+
```agda
707+
recompute-constant : (a? : Dec A) (p q : A) → recompute a? p ≡ recompute a? q
708+
```
709+
693710
* Added new proof in `Relation.Nullary.Decidable`:
694711
```agda
695712
⌊⌋-map′ : (a? : Dec A) → ⌊ map′ t f a? ⌋ ≡ ⌊ a? ⌋
696713
```
697714

715+
* Added new definitions in `Relation.Nullary.Negation.Core`:
716+
```agda
717+
contradiction-irr : .A → ¬ A → Whatever
718+
```
719+
720+
* Added new definitions in `Relation.Nullary.Reflects`:
721+
```agda
722+
recompute : Reflects A b → Recomputable A
723+
recompute-constant : (r : Reflects A b) (p q : A) → recompute r p ≡ recompute r q
724+
```
725+
698726
* Added new definitions in `Relation.Unary`
699727
```
700728
Stable : Pred A ℓ → Set _

src/Axiom/UniquenessOfIdentityProofs.agda

+17-18
Original file line numberDiff line numberDiff line change
@@ -8,23 +8,27 @@
88

99
module 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-constant)
1513
open import Relation.Binary.Core
1614
open import Relation.Binary.Definitions
1715
open import Relation.Binary.PropositionalEquality.Core
1816
open 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
2832
UIP A = Irrelevant {A = A} _≡_
2933

3034
------------------------------------------------------------------------
@@ -39,11 +43,11 @@ UIP A = Irrelevant {A = A} _≡_
3943
-- the image of any other proof.
4044

4145
module Constant⇒UIP
42-
{a} {A : Set a} (f : _≡_ {A = A} ⇒ _≡_)
43-
(f-constant : {a b} (p q : ab) f p ≡ f q)
46+
(f : _≡_ {A = A} ⇒ _≡_)
47+
(f-constant : {x y} (p q : xy) f p ≡ f q)
4448
where
4549

46-
≡-canonical : {a b} (p : ab) trans (sym (f refl)) (f p) ≡ p
50+
≡-canonical : (p : xy) trans (sym (f refl)) (f p) ≡ p
4751
≡-canonical refl = trans-symˡ (f refl)
4852

4953
≡-irrelevant : UIP A
@@ -59,19 +63,14 @@ 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} (_≟_ : DecidableEquality A)
66+
module Decidable⇒UIP (_≟_ : DecidableEquality A)
6467
where
6568

6669
≡-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)
70+
≡-normalise {x} {y} x≡y = recompute (x ≟ y) x≡y
71+
72+
≡-normalise-constant : (p q : x ≡ y) ≡-normalise p ≡ ≡-normalise q
73+
≡-normalise-constant {x = x} {y = y} = recompute-constant (x ≟ y)
7574

7675
≡-irrelevant : UIP A
7776
≡-irrelevant = Constant⇒UIP.≡-irrelevant ≡-normalise ≡-normalise-constant

src/Data/Empty.agda

+3
Original file line numberDiff line numberDiff line change
@@ -35,3 +35,6 @@ private
3535

3636
⊥-elim : {w} {Whatever : Set w} Whatever
3737
⊥-elim ()
38+
39+
⊥-elim-irr : {w} {Whatever : Set w} .⊥ Whatever
40+
⊥-elim-irr ()

src/Relation/Binary/Definitions.agda

+1-3
Original file line numberDiff line numberDiff line change
@@ -12,8 +12,6 @@ module Relation.Binary.Definitions where
1212

1313
open import Agda.Builtin.Equality using (_≡_)
1414

15-
open import Data.Empty using (⊥)
16-
open import Data.Maybe.Base using (Maybe)
1715
open import Data.Product.Base using (_×_; ∃-syntax)
1816
open import Data.Sum.Base using (_⊎_)
1917
open import Function.Base using (_on_; flip)
@@ -248,7 +246,7 @@ Universal _∼_ = ∀ x y → x ∼ y
248246
-- Empty - no elements are related
249247

250248
Empty : REL A B ℓ Set _
251-
Empty _∼_ = {x y} x ∼ y
249+
Empty _∼_ = {x y} ¬ (x ∼ y)
252250

253251
-- Non-emptiness - at least one pair of elements are related.
254252

src/Relation/Nullary.agda

+2-8
Original file line numberDiff line numberDiff line change
@@ -23,8 +23,9 @@ private
2323
------------------------------------------------------------------------
2424
-- Re-exports
2525

26+
open import Relation.Nullary.Recomputable public using (Recomputable)
2627
open import Relation.Nullary.Negation.Core public
27-
open import Relation.Nullary.Reflects public
28+
open import Relation.Nullary.Reflects public hiding (recompute; recompute-constant)
2829
open import Relation.Nullary.Decidable.Core public
2930

3031
------------------------------------------------------------------------
@@ -33,13 +34,6 @@ open import Relation.Nullary.Decidable.Core public
3334
Irrelevant : Set p Set p
3435
Irrelevant P = (p₁ p₂ : P) p₁ ≡ p₂
3536

36-
------------------------------------------------------------------------
37-
-- Recomputability - we can rebuild a relevant proof given an
38-
-- irrelevant one.
39-
40-
Recomputable : Set p Set p
41-
Recomputable P = .P P
42-
4337
------------------------------------------------------------------------
4438
-- Weak decidability
4539
-- `nothing` is 'don't know'/'give up'; `just` is `yes`/`definitely`

src/Relation/Nullary/Decidable.agda

+10-14
Original file line numberDiff line numberDiff line change
@@ -9,15 +9,14 @@
99
module Relation.Nullary.Decidable where
1010

1111
open import Level using (Level)
12-
open import Data.Bool.Base using (true; false; if_then_else_)
13-
open import Data.Empty using (⊥-elim)
12+
open import Data.Bool.Base using (true; false)
1413
open import Data.Product.Base using (∃; _,_)
15-
open import Function.Base
1614
open import Function.Bundles using
1715
(Injection; module Injection; module Equivalence; _⇔_; _↔_; mk↔ₛ′)
1816
open import Relation.Binary.Bundles using (Setoid; module Setoid)
1917
open import Relation.Binary.Definitions using (Decidable)
20-
open import Relation.Nullary
18+
open import Relation.Nullary using (Irrelevant)
19+
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
2120
open import Relation.Nullary.Reflects using (invert)
2221
open import Relation.Binary.PropositionalEquality.Core
2322
using (_≡_; refl; sym; trans; cong′)
@@ -52,8 +51,8 @@ via-injection inj _≟_ x y = map′ injective cong (to x ≟ to y)
5251
-- A lemma relating True and Dec
5352

5453
True-↔ : (a? : Dec A) Irrelevant A True a? ↔ A
55-
True-↔ (true because [a]) irr = mk↔ₛ′ (λ _ invert [a]) _ (irr (invert [a])) cong′
56-
True-↔ (false because ofⁿ ¬a) _ = mk↔ₛ′ (λ ()) (invert (ofⁿ ¬a)) (⊥-elim ∘ ¬a) λ ()
54+
True-↔ (true because [a]) irr = let a = invert [a] in mk↔ₛ′ (λ _ a) _ (irr a) cong′
55+
True-↔ (false because [¬a]) _ = let ¬a = invert [¬a] in mk↔ₛ′ (λ ()) ¬a (λ a contradiction a ¬a) λ ()
5756

5857
------------------------------------------------------------------------
5958
-- Result of decidability
@@ -64,23 +63,20 @@ isYes≗does (false because _) = refl
6463

6564
dec-true : (a? : Dec A) A does a? ≡ true
6665
dec-true (true because _ ) a = refl
67-
dec-true (false because [¬a]) a = ⊥-elim (invert [¬a] a)
66+
dec-true (false because [¬a]) a = contradiction a (invert [¬a])
6867

6968
dec-false : (a? : Dec A) ¬ A does a? ≡ false
7069
dec-false (false because _ ) ¬a = refl
71-
dec-false (true because [a]) ¬a = ⊥-elim (¬a (invert [a]))
70+
dec-false (true because [a]) ¬a = contradiction (invert [a]) ¬a
7271

7372
dec-yes : (a? : Dec A) A λ a a? ≡ yes a
74-
dec-yes a? a with dec-true a? a
75-
dec-yes (yes a′) a | refl = a′ , refl
73+
dec-yes a? a with yes a′ a? | refl dec-true a? a = a′ , refl
7674

7775
dec-no : (a? : Dec A) (¬a : ¬ A) a? ≡ no ¬a
78-
dec-no a? ¬a with dec-false a? ¬a
79-
dec-no (no _) _ | refl = refl
76+
dec-no a? ¬a with no _ a? | refl dec-false a? ¬a = refl
8077

8178
dec-yes-irr : (a? : Dec A) Irrelevant A (a : A) a? ≡ yes a
82-
dec-yes-irr a? irr a with dec-yes a? a
83-
... | a′ , eq rewrite irr a a′ = eq
79+
dec-yes-irr a? irr a with a′ , eq dec-yes a? a rewrite irr a a′ = eq
8480

8581
⌊⌋-map′ : t f (a? : Dec A) ⌊ map′ {B = B} t f a? ⌋ ≡ ⌊ a? ⌋
8682
⌊⌋-map′ t f a? = trans (isYes≗does (map′ t f a?)) (sym (isYes≗does a?))

src/Relation/Nullary/Decidable/Core.agda

+14-9
Original file line numberDiff line numberDiff line change
@@ -11,17 +11,19 @@
1111

1212
module Relation.Nullary.Decidable.Core where
1313

14+
open import Agda.Builtin.Equality using (_≡_)
1415
open import Level using (Level; Lift)
1516
open import Data.Bool.Base using (Bool; T; false; true; not; _∧_; _∨_)
1617
open import Data.Unit.Polymorphic.Base using (⊤)
17-
open import Data.Empty.Irrelevant using (⊥-elim)
1818
open import Data.Product.Base using (_×_)
1919
open import Data.Sum.Base using (_⊎_)
2020
open import Function.Base using (_∘_; const; _$_; flip)
21-
open import Relation.Nullary.Reflects
21+
open import Relation.Nullary.Recomputable as Recomputable hiding (recompute-constant)
22+
open import Relation.Nullary.Reflects as Reflects hiding (recompute; recompute-constant)
2223
open import Relation.Nullary.Negation.Core
2324
using (¬_; Stable; negated-stable; contradiction; DoubleNegation)
2425

26+
2527
private
2628
variable
2729
a b : Level
@@ -69,9 +71,12 @@ module _ {A : Set a} where
6971

7072
-- Given an irrelevant proof of a decidable type, a proof can
7173
-- be recomputed and subsequently used in relevant contexts.
72-
recompute : Dec A .A A
73-
recompute (yes a) _ = a
74-
recompute (no ¬a) a = ⊥-elim (¬a a)
74+
75+
recompute : Dec A Recomputable A
76+
recompute = Reflects.recompute ∘ proof
77+
78+
recompute-constant : (a? : Dec A) (p q : A) recompute a? p ≡ recompute a? q
79+
recompute-constant = Recomputable.recompute-constant ∘ recompute
7580

7681
------------------------------------------------------------------------
7782
-- Interaction with negation, sum, product etc.
@@ -161,17 +166,17 @@ from-no (true because _ ) = _
161166

162167
map′ : (A B) (B A) Dec A Dec B
163168
does (map′ A→B B→A a?) = does a?
164-
proof (map′ A→B B→A (true because [a])) = ofʸ (A→B (invert [a]))
165-
proof (map′ A→B B→A (false because [¬a])) = ofⁿ (invert [¬a] ∘ B→A)
169+
proof (map′ A→B B→A (true because [a])) = of (A→B (invert [a]))
170+
proof (map′ A→B B→A (false because [¬a])) = of (invert [¬a] ∘ B→A)
166171

167172
------------------------------------------------------------------------
168173
-- Relationship with double-negation
169174

170175
-- Decidable predicates are stable.
171176

172177
decidable-stable : Dec A Stable A
173-
decidable-stable (yes a) ¬¬a = a
174-
decidable-stable (no ¬a) ¬¬a = contradiction ¬a ¬¬a
178+
decidable-stable (true because [a]) ¬¬a = invert [a]
179+
decidable-stable (false because [¬a]) ¬¬a = contradiction (invert [¬a]) ¬¬a
175180

176181
¬-drop-Dec : Dec (¬ ¬ A) Dec (¬ A)
177182
¬-drop-Dec ¬¬a? = map′ negated-stable contradiction (¬? ¬¬a?)

0 commit comments

Comments
 (0)