Skip to content

Commit 7a24897

Browse files
committed
Add some conversions
1 parent 4af3db1 commit 7a24897

File tree

4 files changed

+107
-14
lines changed

4 files changed

+107
-14
lines changed

src/Function/Bijection/Strict.agda

+33-2
Original file line numberDiff line numberDiff line change
@@ -15,8 +15,16 @@ open import Relation.Binary.PropositionalEquality as P
1515
open import Function.Equality as F
1616
using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_)
1717
open import Function.Injection as Inj hiding (id; _∘_; injection)
18-
open import Function.Surjection.Strict as Surj hiding (id; _∘_; surjection)
19-
open import Function.LeftInverse.Strict as Left hiding (id; _∘_; leftInverse)
18+
open import Function.Surjection.Strict as Surj hiding (id; _∘_;
19+
surjection;
20+
non-strict-to-strict;
21+
strict-to-non-strict)
22+
open import Function.LeftInverse.Strict as Left hiding (id; _∘_;
23+
leftInverse;
24+
non-strict-to-strict;
25+
strict-to-non-strict)
26+
27+
import Function.Bijection as NonStrict
2028

2129
------------------------------------------------------------------------
2230
-- Bijective functions.
@@ -34,6 +42,29 @@ record Bijective {f₁ f₂ t₁ t₂}
3442
left-inverse-of : from LeftInverseOf to
3543
left-inverse-of x y y≈tox = injective (right-inverse-of (to ⟨$⟩ x) (from ⟨$⟩ y) (F.cong from y≈tox))
3644

45+
------------------------------------------------------------------------
46+
-- Relation to non-strict versions
47+
48+
non-strict-to-strict :
49+
{f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
50+
(f : To ⟶ From)
51+
NonStrict.Bijective f
52+
Bijective f
53+
non-strict-to-strict f ns = record
54+
{ injective = injective
55+
; surjective = Surj.non-strict-to-strict f surjective
56+
} where open NonStrict.Bijective ns
57+
58+
strict-to-non-strict :
59+
{f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
60+
(f : From ⟶ To)
61+
Bijective f
62+
NonStrict.Bijective f
63+
strict-to-non-strict f ns = record
64+
{ injective = injective
65+
; surjective = Surj.strict-to-non-strict f surjective
66+
} where open Bijective ns
67+
3768
------------------------------------------------------------------------
3869
-- The set of all bijections between two setoids.
3970

src/Function/Inverse/Strict.agda

+33-2
Original file line numberDiff line numberDiff line change
@@ -10,14 +10,20 @@ module Function.Inverse.Strict where
1010

1111
open import Level
1212
open import Function.Base using (flip)
13-
open import Function.Bijection.Strict hiding (id; _∘_; bijection)
13+
open import Function.Bijection.Strict hiding (id; _∘_; bijection;
14+
non-strict-to-strict;
15+
strict-to-non-strict)
1416
open import Function.Equality as F
1517
using (_⟶_) renaming (_∘_ to _⟪∘⟫_)
16-
open import Function.LeftInverse.Strict as Left hiding (id; _∘_)
18+
open import Function.LeftInverse.Strict as Left hiding (id; _∘_;
19+
non-strict-to-strict;
20+
strict-to-non-strict)
1721
open import Relation.Binary
1822
open import Relation.Binary.PropositionalEquality as P using (_≗_; _≡_)
1923
open import Relation.Unary using (Pred)
2024

25+
import Function.Inverse as NonStrict
26+
2127
------------------------------------------------------------------------
2228
-- Inverses
2329

@@ -29,6 +35,31 @@ record _InverseOf_ {f₁ f₂ t₁ t₂}
2935
left-inverse-of : from LeftInverseOf to
3036
right-inverse-of : from RightInverseOf to
3137

38+
------------------------------------------------------------------------
39+
-- Relation to non-strict versions
40+
41+
non-strict-to-strict :
42+
{f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
43+
(f : To ⟶ From)
44+
(g : From ⟶ To)
45+
f NonStrict.InverseOf g
46+
f InverseOf g
47+
non-strict-to-strict f g ns = record
48+
{ left-inverse-of = Left.non-strict-to-strict f g left-inverse-of
49+
; right-inverse-of = Left.non-strict-to-strict g f right-inverse-of
50+
} where open NonStrict._InverseOf_ ns
51+
52+
strict-to-non-strict :
53+
{f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
54+
(f : To ⟶ From)
55+
(g : From ⟶ To)
56+
f InverseOf g
57+
f NonStrict.InverseOf g
58+
strict-to-non-strict f g ns = record
59+
{ left-inverse-of = Left.strict-to-non-strict f g left-inverse-of
60+
; right-inverse-of = Left.strict-to-non-strict g f right-inverse-of
61+
} where open _InverseOf_ ns
62+
3263
------------------------------------------------------------------------
3364
-- The set of all inverses between two setoids
3465

src/Function/LeftInverse/Strict.agda

+13-9
Original file line numberDiff line numberDiff line change
@@ -18,7 +18,7 @@ open import Function.Equivalence using (Equivalence)
1818
open import Function.Injection using (Injective; Injection)
1919
open import Relation.Binary.PropositionalEquality as P using (_≡_)
2020

21-
import Function.LeftInverse as LI
21+
import Function.LeftInverse as NonStrict
2222

2323
------------------------------------------------------------------------
2424
-- Left and right inverses.
@@ -41,19 +41,23 @@ f RightInverseOf g = g LeftInverseOf f
4141

4242
non-strict-to-strict :
4343
{f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
44-
LI._LeftInverseOf_ {From = From} {To = To} ⇒
45-
_LeftInverseOf_ {From = From} {To = To}
46-
non-strict-to-strict {From = From} {x = f} {y = g} ns x y y∼gx = begin
44+
(f : To ⟶ From)
45+
(g : From ⟶ To)
46+
f NonStrict.LeftInverseOf g
47+
f LeftInverseOf g
48+
non-strict-to-strict {From = From} f g ns x y y∼gx = begin
4749
f ⟨$⟩ y ≈⟨ Eq.cong f y∼gx ⟩
4850
f ⟨$⟩ (g ⟨$⟩ x) ≈⟨ ns x ⟩
4951
x ∎
5052
where open EqReasoning From
5153

5254
strict-to-non-strict :
5355
{f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
54-
_LeftInverseOf_ {From = From} {To = To} ⇒
55-
LI._LeftInverseOf_ {From = From} {To = To}
56-
strict-to-non-strict {To = To} {x = f} {y = g} str x = str x (g ⟨$⟩ x) refl
56+
(f : To ⟶ From)
57+
(g : From ⟶ To)
58+
f LeftInverseOf g
59+
f NonStrict.LeftInverseOf g
60+
strict-to-non-strict {To = To} f g str x = str x (g ⟨$⟩ x) refl
5761
where open Setoid To
5862

5963
------------------------------------------------------------------------
@@ -70,8 +74,8 @@ record LeftInverse {f₁ f₂ t₁ t₂}
7074
private
7175
open module F = Setoid From
7276
open module T = Setoid To
73-
non-strict : from LI.LeftInverseOf to
74-
non-strict = strict-to-non-strict {x = from} {y = to} left-inverse-of
77+
non-strict : from NonStrict.LeftInverseOf to
78+
non-strict = strict-to-non-strict from to left-inverse-of
7579
open EqReasoning From
7680

7781
injective : Injective to

src/Function/Surjection/Strict.agda

+28-1
Original file line numberDiff line numberDiff line change
@@ -13,11 +13,15 @@ open import Function.Equality as F
1313
using (_⟶_) renaming (_∘_ to _⟪∘⟫_)
1414
open import Function.Equivalence using (Equivalence)
1515
open import Function.Injection hiding (id; _∘_; injection)
16-
open import Function.LeftInverse.Strict as Left hiding (id; _∘_)
16+
open import Function.LeftInverse.Strict as Left hiding (id; _∘_;
17+
non-strict-to-strict;
18+
strict-to-non-strict)
1719
open import Data.Product
1820
open import Relation.Binary
1921
open import Relation.Binary.PropositionalEquality as P using (_≡_)
2022

23+
import Function.Surjection as NonStrict
24+
2125
------------------------------------------------------------------------
2226
-- Surjective functions.
2327

@@ -29,6 +33,29 @@ record Surjective {f₁ f₂ t₁ t₂}
2933
from : To ⟶ From
3034
right-inverse-of : from RightInverseOf to
3135

36+
------------------------------------------------------------------------
37+
-- Relation to non-strict versions
38+
39+
non-strict-to-strict :
40+
{f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
41+
(f : To ⟶ From)
42+
NonStrict.Surjective f
43+
Surjective f
44+
non-strict-to-strict f ns = record
45+
{ from = from
46+
; right-inverse-of = Left.non-strict-to-strict f from right-inverse-of
47+
} where open NonStrict.Surjective ns
48+
49+
strict-to-non-strict :
50+
{f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂}
51+
(f : From ⟶ To)
52+
Surjective f
53+
NonStrict.Surjective f
54+
strict-to-non-strict f ns = record
55+
{ from = from
56+
; right-inverse-of = Left.strict-to-non-strict f from right-inverse-of
57+
} where open Surjective ns
58+
3259
------------------------------------------------------------------------
3360
-- The set of all surjections from one setoid to another.
3461

0 commit comments

Comments
 (0)