@@ -19,7 +19,7 @@ open import Relation.Nullary.Irrelevant using (Irrelevant)
1919open import Relation.Nullary.Negation.Core using (¬_; contradiction)
2020open import Relation.Nullary.Reflects using (invert)
2121open import Relation.Binary.PropositionalEquality.Core
22- using (_≡_; refl; sym; trans; cong′ )
22+ using (_≡_; refl; sym; trans)
2323
2424private
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
5355True-↔ : (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