@@ -19,7 +19,7 @@ open import Relation.Nullary.Irrelevant using (Irrelevant)
19
19
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
20
20
open import Relation.Nullary.Reflects using (invert)
21
21
open import Relation.Binary.PropositionalEquality.Core
22
- using (_≡_; refl; sym; trans; cong′ )
22
+ using (_≡_; refl; sym; trans)
23
23
24
24
private
25
25
variable
@@ -41,18 +41,26 @@ map A⇔B = map′ to from
41
41
-- If there is an injection from one setoid to another, and the
42
42
-- latter's equivalence relation is decidable, then the former's
43
43
-- 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)
49
51
50
52
------------------------------------------------------------------------
51
53
-- A lemma relating True and Dec
52
54
53
55
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
56
64
57
65
------------------------------------------------------------------------
58
66
-- Result of decidability
0 commit comments