-
Notifications
You must be signed in to change notification settings - Fork 246
/
Copy pathInverse.agda
159 lines (125 loc) · 5.17 KB
/
Inverse.agda
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
------------------------------------------------------------------------
-- The Agda standard library
--
-- Properties of inverses.
------------------------------------------------------------------------
{-# OPTIONS --cubical-compatible --safe #-}
module Function.Properties.Inverse where
open import Axiom.Extensionality.Propositional using (Extensionality)
open import Data.Product.Base using (_,_; proj₁; proj₂)
open import Function.Bundles
import Function.Properties.RightInverse as RightInverse
open import Level using (Level; _⊔_)
open import Relation.Binary.Core using (REL)
open import Relation.Binary.Bundles using (Setoid)
open import Relation.Binary.Structures using (IsEquivalence)
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
import Relation.Binary.PropositionalEquality.Properties as P
import Relation.Binary.Reasoning.Setoid as SetoidReasoning
import Function.Consequences.Setoid as Consequences
import Function.Construct.Identity as Identity
import Function.Construct.Symmetry as Symmetry
import Function.Construct.Composition as Composition
private
variable
a b ℓ ℓ₁ ℓ₂ : Level
A B C D : Set a
S T U V : Setoid a ℓ
------------------------------------------------------------------------
-- Setoid bundles
open Identity public using () renaming (inverse to refl)
open Symmetry public using () renaming (inverse to sym)
open Composition public using () renaming (inverse to trans)
isEquivalence : IsEquivalence (Inverse {a} {b})
isEquivalence = record
{ refl = λ {x} → Identity.inverse x
; sym = sym
; trans = trans
}
------------------------------------------------------------------------
-- Propositional bundles
↔-refl : A ↔ A
↔-refl = Identity.↔-id _
↔-sym : A ↔ B → B ↔ A
↔-sym = Symmetry.↔-sym
↔-trans : A ↔ B → B ↔ C → A ↔ C
↔-trans = Composition.inverse
-- need to η-expand for everything to line up properly
↔-isEquivalence : IsEquivalence {ℓ = ℓ} _↔_
↔-isEquivalence = record
{ refl = ↔-refl
; sym = ↔-sym
; trans = ↔-trans
}
------------------------------------------------------------------------
-- Conversion functions
toFunction : Inverse S T → Func S T
toFunction I = record { to = to ; cong = to-cong }
where open Inverse I
fromFunction : Inverse S T → Func T S
fromFunction I = record { to = from ; cong = from-cong }
where open Inverse I
Inverse⇒Injection : Inverse S T → Injection S T
Inverse⇒Injection {S = S} {T = T} I = record
{ to = to
; cong = to-cong
; injective = inverseʳ⇒injective to inverseʳ
} where open Inverse I; open Consequences S T
Inverse⇒Surjection : Inverse S T → Surjection S T
Inverse⇒Surjection {S = S} {T = T} I = record
{ to = to
; cong = to-cong
; surjective = inverseˡ⇒surjective inverseˡ
} where open Inverse I; open Consequences S T
Inverse⇒Bijection : Inverse S T → Bijection S T
Inverse⇒Bijection {S = S} {T = T} I = record
{ to = to
; cong = to-cong
; bijective = inverseᵇ⇒bijective inverse
} where open Inverse I; open Consequences S T
Inverse⇒Equivalence : Inverse S T → Equivalence S T
Inverse⇒Equivalence I = record
{ to = to
; from = from
; to-cong = to-cong
; from-cong = from-cong
} where open Inverse I
↔⇒⟶ : A ↔ B → A ⟶ B
↔⇒⟶ = toFunction
↔⇒⟵ : A ↔ B → B ⟶ A
↔⇒⟵ = fromFunction
↔⇒↣ : A ↔ B → A ↣ B
↔⇒↣ = Inverse⇒Injection
↔⇒↠ : A ↔ B → A ↠ B
↔⇒↠ = Inverse⇒Surjection
↔⇒⤖ : A ↔ B → A ⤖ B
↔⇒⤖ = Inverse⇒Bijection
↔⇒⇔ : A ↔ B → A ⇔ B
↔⇒⇔ = Inverse⇒Equivalence
↔⇒↩ : A ↔ B → A ↩ B
↔⇒↩ = Inverse.leftInverse
↔⇒↪ : A ↔ B → A ↪ B
↔⇒↪ = Inverse.rightInverse
-- The functions above can be combined with the following lemma to
-- transport an arbitrary relation R (e.g. Injection) across
-- inverses.
transportVia : {R : ∀ {a b ℓ₁ ℓ₂} → REL (Setoid a ℓ₁) (Setoid b ℓ₂) (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂)} →
(∀ {a b c ℓ₁ ℓ₂ ℓ₃} {S : Setoid a ℓ₁} {T : Setoid b ℓ₂} {U : Setoid c ℓ₃} → R S T → R T U → R S U) →
(∀ {a b ℓ₁ ℓ₂} {S : Setoid a ℓ₁} {T : Setoid b ℓ₂} → Inverse S T → R S T) →
Inverse S T → R T U → Inverse U V → R S V
transportVia R-trans inv⇒R IBA RBC ICD =
R-trans (inv⇒R IBA) (R-trans RBC (inv⇒R ICD))
------------------------------------------------------------------------
-- Other
module _ (ext : ∀ {a b} → Extensionality a b) where
↔-fun : A ↔ B → C ↔ D → (A → C) ↔ (B → D)
↔-fun A↔B C↔D = mk↔ₛ′
(λ a→c b → to C↔D (a→c (from A↔B b)))
(λ b→d a → from C↔D (b→d (to A↔B a)))
(λ b→d → ext λ _ → P.trans (strictlyInverseˡ C↔D _ ) (P.cong b→d (strictlyInverseˡ A↔B _)))
(λ a→c → ext λ _ → P.trans (strictlyInverseʳ C↔D _ ) (P.cong a→c (strictlyInverseʳ A↔B _)))
where open Inverse
module _ (I : Inverse S T) where
open Inverse I
to-from : ∀ {x y} → to x Eq₂.≈ y → from y Eq₁.≈ x
to-from = RightInverse.to-from rightInverse