Skip to content

Commit cf9c04e

Browse files
committed
fix: knock-ons
1 parent c15f748 commit cf9c04e

File tree

3 files changed

+48
-41
lines changed

3 files changed

+48
-41
lines changed

src/Data/Product/Function/Dependent/Propositional.agda

Lines changed: 13 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -194,17 +194,21 @@ module _ where
194194
to′ = map (to I↠J) (to A↠B)
195195

196196
backcast : {i} B i B (to I↠J (section I↠J i))
197-
backcast = ≡.subst B (≡.sym (to∘to⁻ I↠J _))
197+
backcast = ≡.subst B (≡.sym (section-strictInverseˡ I↠J _))
198198

199-
to⁻: Σ J B Σ I A
200-
to⁻= map (section I↠J) (Surjection.section A↠B ∘ backcast)
199+
section: Σ J B Σ I A
200+
section= map (section I↠J) (section A↠B ∘ backcast)
201201

202202
strictlySurjective′ : StrictlySurjective _≡_ to′
203-
strictlySurjective′ (x , y) = to⁻′ (x , y) , Σ-≡,≡→≡
204-
( to∘to⁻ I↠J x
205-
, (≡.subst B (to∘to⁻ I↠J x) (to A↠B (section A↠B (backcast y))) ≡⟨ ≡.cong (≡.subst B _) (to∘to⁻ A↠B _) ⟩
206-
≡.subst B (to∘to⁻ I↠J x) (backcast y) ≡⟨ ≡.subst-subst-sym (to∘to⁻ I↠J x) ⟩
207-
y ∎)
203+
strictlySurjective′ (x , y) = section′ (x , y) , Σ-≡,≡→≡
204+
( section-strictInverseˡ I↠J x
205+
, (begin
206+
≡.subst B (section-strictInverseˡ I↠J x) (to A↠B (section A↠B (backcast y)))
207+
≡⟨ ≡.cong (≡.subst B _) (section-strictInverseˡ A↠B _) ⟩
208+
≡.subst B (section-strictInverseˡ I↠J x) (backcast y)
209+
≡⟨ ≡.subst-subst-sym (section-strictInverseˡ I↠J x) ⟩
210+
y
211+
∎)
208212
) where open ≡.≡-Reasoning
209213

210214

@@ -250,7 +254,7 @@ module _ where
250254
Σ-↔ {I = I} {J = J} {A = A} {B = B} I↔J A↔B = mk↔ₛ′
251255
(Surjection.to surjection′)
252256
(Surjection.section surjection′)
253-
(Surjection.to∘to⁻ surjection′)
257+
(Surjection.section-strictInverseˡ surjection′)
254258
left-inverse-of
255259
where
256260
open ≡.≡-Reasoning

src/Data/Product/Function/Dependent/Setoid.agda

Lines changed: 34 additions & 31 deletions
Original file line numberDiff line numberDiff line change
@@ -94,34 +94,37 @@ module _ where
9494
(function (⇔⇒⟶ I⇔J) A⟶B)
9595
(function (⇔⇒⟵ I⇔J) B⟶A)
9696

97-
equivalence-↪ :
98-
(I↪J : I ↪ J)
99-
( {i} Equivalence (A atₛ (RightInverse.from I↪J i)) (B atₛ i))
100-
Equivalence (I ×ₛ A) (J ×ₛ B)
101-
equivalence-↪ {A = A} {B = B} I↪J A⇔B =
102-
equivalence (RightInverse.equivalence I↪J) A→B (fromFunction A⇔B)
103-
where
104-
A→B : {i} Func (A atₛ i) (B atₛ (RightInverse.to I↪J i))
105-
A→B = record
106-
{ to = to A⇔B ∘ cast A (RightInverse.strictlyInverseʳ I↪J _)
107-
; cong = to-cong A⇔B ∘ cast-cong A (RightInverse.strictlyInverseʳ I↪J _)
108-
}
109-
110-
equivalence-↠ :
111-
(I↠J : I ↠ J)
112-
( {x} Equivalence (A atₛ x) (B atₛ (Surjection.to I↠J x)))
113-
Equivalence (I ×ₛ A) (J ×ₛ B)
114-
equivalence-↠ {A = A} {B = B} I↠J A⇔B =
115-
equivalence (↠⇒⇔ I↠J) B-to B-from
116-
where
117-
B-to : {x} Func (A atₛ x) (B atₛ (Surjection.to I↠J x))
118-
B-to = toFunction A⇔B
97+
module _ (I↪J : I ↪ J) where
98+
99+
private module ItoJ = RightInverse I↪J
100+
101+
equivalence-↪ : ( {i} Equivalence (A atₛ (ItoJ.from i)) (B atₛ i))
102+
Equivalence (I ×ₛ A) (J ×ₛ B)
103+
equivalence-↪ {A = A} {B = B} A⇔B =
104+
equivalence ItoJ.equivalence A→B (fromFunction A⇔B)
105+
where
106+
A→B : {i} Func (A atₛ i) (B atₛ (ItoJ.to i))
107+
A→B = record
108+
{ to = to A⇔B ∘ cast A (ItoJ.strictlyInverseʳ _)
109+
; cong = to-cong A⇔B ∘ cast-cong A (ItoJ.strictlyInverseʳ _)
110+
}
111+
112+
module _ (I↠J : I ↠ J) where
113+
114+
private module ItoJ = Surjection I↠J
115+
116+
equivalence-↠ : ( {x} Equivalence (A atₛ x) (B atₛ (ItoJ.to x)))
117+
Equivalence (I ×ₛ A) (J ×ₛ B)
118+
equivalence-↠ {A = A} {B = B} A⇔B = equivalence (↠⇒⇔ I↠J) B-to B-from
119+
where
120+
B-to : {x} Func (A atₛ x) (B atₛ (ItoJ.to x))
121+
B-to = toFunction A⇔B
119122

120-
B-from : {y} Func (B atₛ y) (A atₛ (Surjection.section I↠J y))
121-
B-from = record
122-
{ to = from A⇔B ∘ cast B (Surjection.to∘to⁻ I↠J _)
123-
; cong = from-cong A⇔B ∘ cast-cong B (Surjection.to∘to⁻ I↠J _)
124-
}
123+
B-from : {y} Func (B atₛ y) (A atₛ (ItoJ.section y))
124+
B-from = record
125+
{ to = from A⇔B ∘ cast B (ItoJ.section-strictInverseˡ _)
126+
; cong = from-cong A⇔B ∘ cast-cong B (ItoJ.section-strictInverseˡ _)
127+
}
125128

126129
------------------------------------------------------------------------
127130
-- Injections
@@ -168,12 +171,12 @@ module _ where
168171
func : Func (I ×ₛ A) (J ×ₛ B)
169172
func = function (Surjection.function I↠J) (Surjection.function A↠B)
170173

171-
to⁻: Carrier (J ×ₛ B) Carrier (I ×ₛ A)
172-
to⁻′ (j , y) = section I↠J j , section A↠B (cast B (Surjection.to∘to⁻ I↠J _) y)
174+
section: Carrier (J ×ₛ B) Carrier (I ×ₛ A)
175+
section′ (j , y) = section I↠J j , section A↠B (cast B (section-strictInverseˡ I↠J _) y)
173176

174177
strictlySurj : StrictlySurjective (Func.Eq₂._≈_ func) (Func.to func)
175-
strictlySurj (j , y) = to⁻′ (j , y) ,
176-
to∘to⁻ I↠J j , IndexedSetoid.trans B (to∘to⁻ A↠B _) (cast-eq B (to∘to⁻ I↠J j))
178+
strictlySurj (j , y) = section′ (j , y) ,
179+
section-strictInverseˡ I↠J j , IndexedSetoid.trans B (section-strictInverseˡ A↠B _) (cast-eq B (section-strictInverseˡ I↠J j))
177180

178181
surj : Surjective (Func.Eq₁._≈_ func) (Func.Eq₂._≈_ func) (Func.to func)
179182
surj = strictlySurjective⇒surjective (I ×ₛ A) (J ×ₛ B) (Func.cong func) strictlySurj

src/Function/Properties/Surjection.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -45,7 +45,7 @@ mkSurjection f surjective = record
4545
↠⇒⟶ = Surjection.function
4646

4747
↠⇒↪ : A ↠ B B ↪ A
48-
↠⇒↪ s = mk↪ {from = to} λ { ≡.refl to∘to⁻ _ }
48+
↠⇒↪ s = mk↪ {from = to} λ { ≡.refl section-strictInverseˡ _ }
4949
where open Surjection s
5050

5151
↠⇒⇔ : A ↠ B A ⇔ B

0 commit comments

Comments
 (0)