Skip to content

Commit 32f786e

Browse files
authored
[ refactor ] proofs in Relation.Nullary.Decidable and a reexport in Function.Related.TypeIsomorphisms (#2737)
* refactor: one proof, and its reexport * refactor: one more proof * refactor: streamline * fix: `import`s
1 parent 2618431 commit 32f786e

File tree

2 files changed

+24
-21
lines changed

2 files changed

+24
-21
lines changed

src/Function/Related/TypeIsomorphisms.agda

Lines changed: 8 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -37,16 +37,21 @@ open import Relation.Binary hiding (_⇔_)
3737
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong)
3838
open import Relation.Binary.PropositionalEquality.Properties
3939
using (module ≡-Reasoning)
40-
open import Relation.Nullary.Reflects using (invert)
41-
open import Relation.Nullary using (Dec; ¬_; _because_; ofⁿ; contradiction)
40+
open import Relation.Nullary.Negation.Core using (¬_)
4241
import Relation.Nullary.Indexed as I
43-
open import Relation.Nullary.Decidable using (True)
4442

4543
private
4644
variable
4745
a b c d : Level
4846
A B C D : Set a
4947

48+
------------------------------------------------------------------------
49+
-- A lemma relating True dec and P, where dec : Dec P
50+
51+
open import Relation.Nullary.Decidable public
52+
using ()
53+
renaming (True-↔ to True↔)
54+
5055
------------------------------------------------------------------------
5156
-- Properties of Σ and _×_
5257

@@ -350,16 +355,6 @@ Related-cong {A = A} {B = B} {C = C} {D = D} A≈B C≈D = mk⇔
350355
C ∎)
351356
where open EquationalReasoning
352357

353-
------------------------------------------------------------------------
354-
-- A lemma relating True dec and P, where dec : Dec P
355-
356-
True↔ : {p} {P : Set p}
357-
(dec : Dec P) ((p₁ p₂ : P) p₁ ≡ p₂) True dec ↔ P
358-
True↔ ( true because [p]) irr =
359-
mk↔ₛ′ (λ _ invert [p]) (λ _ _) (irr _) (λ _ refl)
360-
True↔ (false because ofⁿ ¬p) _ =
361-
mk↔ₛ′ (λ()) (invert (ofⁿ ¬p)) (λ x flip contradiction ¬p x) (λ ())
362-
363358
------------------------------------------------------------------------
364359
-- Relating a predicate to an existentially quantified one with the
365360
-- restriction that the quantified variable is equal to the given one

src/Relation/Nullary/Decidable.agda

Lines changed: 16 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -19,7 +19,7 @@ open import Relation.Nullary.Irrelevant using (Irrelevant)
1919
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
2020
open import Relation.Nullary.Reflects using (invert)
2121
open import Relation.Binary.PropositionalEquality.Core
22-
using (_≡_; refl; sym; trans; cong′)
22+
using (_≡_; refl; sym; trans)
2323

2424
private
2525
variable
@@ -41,18 +41,26 @@ map A⇔B = map′ to from
4141
-- If there is an injection from one setoid to another, and the
4242
-- latter's equivalence relation is decidable, then the former's
4343
-- equivalence relation is also decidable.
44-
via-injection : {S : Setoid a ℓ₁} {T : Setoid b ℓ₂}
45-
(inj : Injection S T) (open Injection inj)
46-
Decidable Eq₂._≈_ Decidable Eq₁._≈_
47-
via-injection inj _≟_ x y = map′ injective cong (to x ≟ to y)
48-
where open Injection inj
44+
45+
module _ {S : Setoid a ℓ₁} {T : Setoid b ℓ₂} (injection : Injection S T) where
46+
47+
open Injection injection
48+
49+
via-injection : Decidable Eq₂._≈_ Decidable Eq₁._≈_
50+
via-injection _≟_ x y = map′ injective cong (to x ≟ to y)
4951

5052
------------------------------------------------------------------------
5153
-- A lemma relating True and Dec
5254

5355
True-↔ : (a? : Dec A) Irrelevant A True a? ↔ 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) λ ()
56+
True-↔ a? irr = mk↔ₛ′ to from to-from (from-to a?)
57+
where
58+
to = toWitness {a? = a?}
59+
from = fromWitness {a? = a?}
60+
to-from : a to (from a) ≡ a
61+
to-from a = irr _ a
62+
from-to : a? (x : True a?) fromWitness (toWitness x) ≡ x
63+
from-to (yes _) _ = refl
5664

5765
------------------------------------------------------------------------
5866
-- Result of decidability

0 commit comments

Comments
 (0)