mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
129 lines
No EOL
31 KiB
HTML
129 lines
No EOL
31 KiB
HTML
<!DOCTYPE HTML>
|
||
<html><head><meta charset="utf-8"><title>Function.Structures.Biased</title><link rel="stylesheet" href="Agda.css"></head><body><pre class="Agda"><a id="1" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="74" class="Comment">-- The Agda standard library</a>
|
||
<a id="103" class="Comment">--</a>
|
||
<a id="106" class="Comment">-- Ways to give instances of certain structures where some fields can</a>
|
||
<a id="176" class="Comment">-- be given in terms of others.</a>
|
||
<a id="208" class="Comment">-- The contents of this file should usually be accessed from `Function`.</a>
|
||
<a id="281" class="Comment">------------------------------------------------------------------------</a>
|
||
|
||
|
||
<a id="356" class="Symbol">{-#</a> <a id="360" class="Keyword">OPTIONS</a> <a id="368" class="Pragma">--cubical-compatible</a> <a id="389" class="Pragma">--safe</a> <a id="396" class="Symbol">#-}</a>
|
||
|
||
<a id="401" class="Keyword">open</a> <a id="406" class="Keyword">import</a> <a id="413" href="Relation.Binary.Core.html" class="Module">Relation.Binary.Core</a> <a id="434" class="Keyword">using</a> <a id="440" class="Symbol">(</a><a id="441" href="Relation.Binary.Core.html#896" class="Function">Rel</a><a id="444" class="Symbol">)</a>
|
||
<a id="446" class="Keyword">open</a> <a id="451" class="Keyword">import</a> <a id="458" href="Relation.Binary.Bundles.html" class="Module">Relation.Binary.Bundles</a> <a id="482" class="Keyword">using</a> <a id="488" class="Symbol">(</a><a id="489" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a><a id="495" class="Symbol">)</a>
|
||
<a id="497" class="Keyword">open</a> <a id="502" class="Keyword">import</a> <a id="509" href="Relation.Binary.Structures.html" class="Module">Relation.Binary.Structures</a> <a id="536" class="Keyword">using</a> <a id="542" class="Symbol">(</a><a id="543" href="Relation.Binary.Structures.html#1550" class="Record">IsEquivalence</a><a id="556" class="Symbol">)</a>
|
||
|
||
<a id="559" class="Keyword">module</a> <a id="566" href="Function.Structures.Biased.html" class="Module">Function.Structures.Biased</a> <a id="593" class="Symbol">{</a><a id="594" href="Function.Structures.Biased.html#594" class="Bound">a</a> <a id="596" href="Function.Structures.Biased.html#596" class="Bound">b</a> <a id="598" href="Function.Structures.Biased.html#598" class="Bound">ℓ₁</a> <a id="601" href="Function.Structures.Biased.html#601" class="Bound">ℓ₂</a><a id="603" class="Symbol">}</a>
|
||
<a id="607" class="Symbol">{</a><a id="608" href="Function.Structures.Biased.html#608" class="Bound">A</a> <a id="610" class="Symbol">:</a> <a id="612" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="616" href="Function.Structures.Biased.html#594" class="Bound">a</a><a id="617" class="Symbol">}</a> <a id="619" class="Symbol">(</a><a id="620" href="Function.Structures.Biased.html#620" class="Bound Operator">_≈₁_</a> <a id="625" class="Symbol">:</a> <a id="627" href="Relation.Binary.Core.html#896" class="Function">Rel</a> <a id="631" href="Function.Structures.Biased.html#608" class="Bound">A</a> <a id="633" href="Function.Structures.Biased.html#598" class="Bound">ℓ₁</a><a id="635" class="Symbol">)</a> <a id="637" class="Comment">-- Equality over the domain</a>
|
||
<a id="667" class="Symbol">{</a><a id="668" href="Function.Structures.Biased.html#668" class="Bound">B</a> <a id="670" class="Symbol">:</a> <a id="672" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="676" href="Function.Structures.Biased.html#596" class="Bound">b</a><a id="677" class="Symbol">}</a> <a id="679" class="Symbol">(</a><a id="680" href="Function.Structures.Biased.html#680" class="Bound Operator">_≈₂_</a> <a id="685" class="Symbol">:</a> <a id="687" href="Relation.Binary.Core.html#896" class="Function">Rel</a> <a id="691" href="Function.Structures.Biased.html#668" class="Bound">B</a> <a id="693" href="Function.Structures.Biased.html#601" class="Bound">ℓ₂</a><a id="695" class="Symbol">)</a> <a id="697" class="Comment">-- Equality over the codomain</a>
|
||
<a id="729" class="Keyword">where</a>
|
||
|
||
<a id="736" class="Keyword">open</a> <a id="741" class="Keyword">import</a> <a id="748" href="Data.Product.Base.html" class="Module">Data.Product.Base</a> <a id="766" class="Symbol">as</a> <a id="769" class="Module">Product</a> <a id="777" class="Keyword">using</a> <a id="783" class="Symbol">(</a><a id="784" href="Data.Product.Base.html#852" class="Function">∃</a><a id="785" class="Symbol">;</a> <a id="787" href="Data.Product.Base.html#1618" class="Function Operator">_×_</a><a id="790" class="Symbol">;</a> <a id="792" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">_,_</a><a id="795" class="Symbol">)</a>
|
||
<a id="797" class="Keyword">open</a> <a id="802" class="Keyword">import</a> <a id="809" href="Function.Base.html" class="Module">Function.Base</a>
|
||
<a id="823" class="Keyword">open</a> <a id="828" class="Keyword">import</a> <a id="835" href="Function.Definitions.html" class="Module">Function.Definitions</a>
|
||
<a id="856" class="Keyword">open</a> <a id="861" class="Keyword">import</a> <a id="868" href="Function.Structures.html" class="Module">Function.Structures</a> <a id="888" href="Function.Structures.Biased.html#620" class="Bound Operator">_≈₁_</a> <a id="893" href="Function.Structures.Biased.html#680" class="Bound Operator">_≈₂_</a>
|
||
<a id="898" class="Keyword">open</a> <a id="903" class="Keyword">import</a> <a id="910" href="Function.Consequences.Setoid.html" class="Module">Function.Consequences.Setoid</a>
|
||
<a id="939" class="Keyword">open</a> <a id="944" class="Keyword">import</a> <a id="951" href="Level.html" class="Module">Level</a> <a id="957" class="Keyword">using</a> <a id="963" class="Symbol">(</a><a id="964" href="Agda.Primitive.html#961" class="Primitive Operator">_⊔_</a><a id="967" class="Symbol">)</a>
|
||
|
||
<a id="970" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="1043" class="Comment">-- Surjection</a>
|
||
<a id="1057" class="Comment">------------------------------------------------------------------------</a>
|
||
|
||
<a id="1131" class="Keyword">record</a> <a id="IsStrictSurjection"></a><a id="1138" href="Function.Structures.Biased.html#1138" class="Record">IsStrictSurjection</a> <a id="1157" class="Symbol">(</a><a id="1158" href="Function.Structures.Biased.html#1158" class="Bound">f</a> <a id="1160" class="Symbol">:</a> <a id="1162" href="Function.Structures.Biased.html#608" class="Bound">A</a> <a id="1164" class="Symbol">→</a> <a id="1166" href="Function.Structures.Biased.html#668" class="Bound">B</a><a id="1167" class="Symbol">)</a> <a id="1169" class="Symbol">:</a> <a id="1171" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1175" class="Symbol">(</a><a id="1176" href="Function.Structures.Biased.html#594" class="Bound">a</a> <a id="1178" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1180" href="Function.Structures.Biased.html#596" class="Bound">b</a> <a id="1182" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1184" href="Function.Structures.Biased.html#598" class="Bound">ℓ₁</a> <a id="1187" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1189" href="Function.Structures.Biased.html#601" class="Bound">ℓ₂</a><a id="1191" class="Symbol">)</a> <a id="1193" class="Keyword">where</a>
|
||
<a id="1201" class="Keyword">field</a>
|
||
<a id="IsStrictSurjection.isCongruent"></a><a id="1211" href="Function.Structures.Biased.html#1211" class="Field">isCongruent</a> <a id="1223" class="Symbol">:</a> <a id="1225" href="Function.Structures.html#995" class="Record">IsCongruent</a> <a id="1237" href="Function.Structures.Biased.html#1158" class="Bound">f</a>
|
||
<a id="IsStrictSurjection.strictlySurjective"></a><a id="1243" href="Function.Structures.Biased.html#1243" class="Field">strictlySurjective</a> <a id="1262" class="Symbol">:</a> <a id="1264" href="Function.Definitions.html#1522" class="Function">StrictlySurjective</a> <a id="1283" href="Function.Structures.Biased.html#680" class="Bound Operator">_≈₂_</a> <a id="1288" href="Function.Structures.Biased.html#1158" class="Bound">f</a>
|
||
|
||
<a id="1293" class="Keyword">open</a> <a id="1298" href="Function.Structures.html#995" class="Module">IsCongruent</a> <a id="1310" href="Function.Structures.Biased.html#1211" class="Field">isCongruent</a> <a id="1322" class="Keyword">public</a>
|
||
|
||
<a id="IsStrictSurjection.isSurjection"></a><a id="1332" href="Function.Structures.Biased.html#1332" class="Function">isSurjection</a> <a id="1345" class="Symbol">:</a> <a id="1347" href="Function.Structures.html#1664" class="Record">IsSurjection</a> <a id="1360" href="Function.Structures.Biased.html#1158" class="Bound">f</a>
|
||
<a id="1364" href="Function.Structures.Biased.html#1332" class="Function">isSurjection</a> <a id="1377" class="Symbol">=</a> <a id="1379" class="Keyword">record</a>
|
||
<a id="1390" class="Symbol">{</a> <a id="1392" href="Function.Structures.html#1731" class="Field">isCongruent</a> <a id="1404" class="Symbol">=</a> <a id="1406" href="Function.Structures.Biased.html#1211" class="Field">isCongruent</a>
|
||
<a id="1422" class="Symbol">;</a> <a id="1424" href="Function.Structures.html#1763" class="Field">surjective</a> <a id="1435" class="Symbol">=</a> <a id="1437" href="Function.Consequences.Setoid.html#1959" class="Function">strictlySurjective⇒surjective</a>
|
||
<a id="1475" href="Function.Structures.html#1207" class="Function">Eq₁.setoid</a> <a id="1486" href="Function.Structures.html#1351" class="Function">Eq₂.setoid</a> <a id="1497" href="Function.Structures.html#1062" class="Function">cong</a> <a id="1502" href="Function.Structures.Biased.html#1243" class="Field">strictlySurjective</a>
|
||
<a id="1525" class="Symbol">}</a>
|
||
|
||
<a id="1528" class="Keyword">open</a> <a id="1533" href="Function.Structures.Biased.html#1138" class="Module">IsStrictSurjection</a> <a id="1552" class="Keyword">public</a>
|
||
<a id="1561" class="Keyword">using</a> <a id="1567" class="Symbol">()</a> <a id="1570" class="Keyword">renaming</a> <a id="1579" class="Symbol">(</a><a id="1580" href="Function.Structures.Biased.html#1332" class="Function">isSurjection</a> <a id="1593" class="Symbol">to</a> <a id="1596" class="Function">isStrictSurjection</a><a id="1614" class="Symbol">)</a>
|
||
|
||
<a id="1617" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="1690" class="Comment">-- Bijection</a>
|
||
<a id="1703" class="Comment">------------------------------------------------------------------------</a>
|
||
|
||
<a id="1777" class="Keyword">record</a> <a id="IsStrictBijection"></a><a id="1784" href="Function.Structures.Biased.html#1784" class="Record">IsStrictBijection</a> <a id="1802" class="Symbol">(</a><a id="1803" href="Function.Structures.Biased.html#1803" class="Bound">f</a> <a id="1805" class="Symbol">:</a> <a id="1807" href="Function.Structures.Biased.html#608" class="Bound">A</a> <a id="1809" class="Symbol">→</a> <a id="1811" href="Function.Structures.Biased.html#668" class="Bound">B</a><a id="1812" class="Symbol">)</a> <a id="1814" class="Symbol">:</a> <a id="1816" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1820" class="Symbol">(</a><a id="1821" href="Function.Structures.Biased.html#594" class="Bound">a</a> <a id="1823" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1825" href="Function.Structures.Biased.html#596" class="Bound">b</a> <a id="1827" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1829" href="Function.Structures.Biased.html#598" class="Bound">ℓ₁</a> <a id="1832" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1834" href="Function.Structures.Biased.html#601" class="Bound">ℓ₂</a><a id="1836" class="Symbol">)</a> <a id="1838" class="Keyword">where</a>
|
||
<a id="1846" class="Keyword">field</a>
|
||
<a id="IsStrictBijection.isInjection"></a><a id="1856" href="Function.Structures.Biased.html#1856" class="Field">isInjection</a> <a id="1868" class="Symbol">:</a> <a id="1870" href="Function.Structures.html#1479" class="Record">IsInjection</a> <a id="1882" href="Function.Structures.Biased.html#1803" class="Bound">f</a>
|
||
<a id="IsStrictBijection.strictlySurjective"></a><a id="1888" href="Function.Structures.Biased.html#1888" class="Field">strictlySurjective</a> <a id="1908" class="Symbol">:</a> <a id="1910" href="Function.Definitions.html#1522" class="Function">StrictlySurjective</a> <a id="1929" href="Function.Structures.Biased.html#680" class="Bound Operator">_≈₂_</a> <a id="1934" href="Function.Structures.Biased.html#1803" class="Bound">f</a>
|
||
|
||
<a id="IsStrictBijection.isBijection"></a><a id="1939" href="Function.Structures.Biased.html#1939" class="Function">isBijection</a> <a id="1951" class="Symbol">:</a> <a id="1953" href="Function.Structures.html#1970" class="Record">IsBijection</a> <a id="1965" href="Function.Structures.Biased.html#1803" class="Bound">f</a>
|
||
<a id="1969" href="Function.Structures.Biased.html#1939" class="Function">isBijection</a> <a id="1981" class="Symbol">=</a> <a id="1983" class="Keyword">record</a>
|
||
<a id="1994" class="Symbol">{</a> <a id="1996" href="Function.Structures.html#2036" class="Field">isInjection</a> <a id="2008" class="Symbol">=</a> <a id="2010" href="Function.Structures.Biased.html#1856" class="Field">isInjection</a>
|
||
<a id="2026" class="Symbol">;</a> <a id="2028" href="Function.Structures.html#2068" class="Field">surjective</a> <a id="2039" class="Symbol">=</a> <a id="2041" href="Function.Consequences.Setoid.html#1959" class="Function">strictlySurjective⇒surjective</a>
|
||
<a id="2079" href="Function.Structures.html#1207" class="Function">Eq₁.setoid</a> <a id="2090" href="Function.Structures.html#1351" class="Function">Eq₂.setoid</a> <a id="2101" href="Function.Structures.html#1062" class="Function">cong</a> <a id="2106" href="Function.Structures.Biased.html#1888" class="Field">strictlySurjective</a>
|
||
<a id="2129" class="Symbol">}</a> <a id="2131" class="Keyword">where</a> <a id="2137" class="Keyword">open</a> <a id="2142" href="Function.Structures.html#1479" class="Module">IsInjection</a> <a id="2154" href="Function.Structures.Biased.html#1856" class="Field">isInjection</a>
|
||
|
||
<a id="2167" class="Keyword">open</a> <a id="2172" href="Function.Structures.Biased.html#1784" class="Module">IsStrictBijection</a> <a id="2190" class="Keyword">public</a>
|
||
<a id="2199" class="Keyword">using</a> <a id="2205" class="Symbol">()</a> <a id="2208" class="Keyword">renaming</a> <a id="2217" class="Symbol">(</a><a id="2218" href="Function.Structures.Biased.html#1939" class="Function">isBijection</a> <a id="2230" class="Symbol">to</a> <a id="2233" class="Function">isStrictBijection</a><a id="2250" class="Symbol">)</a>
|
||
|
||
<a id="2253" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="2326" class="Comment">-- Left inverse</a>
|
||
<a id="2342" class="Comment">------------------------------------------------------------------------</a>
|
||
|
||
<a id="2416" class="Keyword">record</a> <a id="IsStrictLeftInverse"></a><a id="2423" href="Function.Structures.Biased.html#2423" class="Record">IsStrictLeftInverse</a> <a id="2443" class="Symbol">(</a><a id="2444" href="Function.Structures.Biased.html#2444" class="Bound">to</a> <a id="2447" class="Symbol">:</a> <a id="2449" href="Function.Structures.Biased.html#608" class="Bound">A</a> <a id="2451" class="Symbol">→</a> <a id="2453" href="Function.Structures.Biased.html#668" class="Bound">B</a><a id="2454" class="Symbol">)</a> <a id="2456" class="Symbol">(</a><a id="2457" href="Function.Structures.Biased.html#2457" class="Bound">from</a> <a id="2462" class="Symbol">:</a> <a id="2464" href="Function.Structures.Biased.html#668" class="Bound">B</a> <a id="2466" class="Symbol">→</a> <a id="2468" href="Function.Structures.Biased.html#608" class="Bound">A</a><a id="2469" class="Symbol">)</a> <a id="2471" class="Symbol">:</a> <a id="2473" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="2477" class="Symbol">(</a><a id="2478" href="Function.Structures.Biased.html#594" class="Bound">a</a> <a id="2480" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2482" href="Function.Structures.Biased.html#596" class="Bound">b</a> <a id="2484" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2486" href="Function.Structures.Biased.html#598" class="Bound">ℓ₁</a> <a id="2489" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2491" href="Function.Structures.Biased.html#601" class="Bound">ℓ₂</a><a id="2493" class="Symbol">)</a> <a id="2495" class="Keyword">where</a>
|
||
<a id="2503" class="Keyword">field</a>
|
||
<a id="IsStrictLeftInverse.isCongruent"></a><a id="2513" href="Function.Structures.Biased.html#2513" class="Field">isCongruent</a> <a id="2526" class="Symbol">:</a> <a id="2528" href="Function.Structures.html#995" class="Record">IsCongruent</a> <a id="2540" href="Function.Structures.Biased.html#2444" class="Bound">to</a>
|
||
<a id="IsStrictLeftInverse.from-cong"></a><a id="2547" href="Function.Structures.Biased.html#2547" class="Field">from-cong</a> <a id="2560" class="Symbol">:</a> <a id="2562" href="Function.Definitions.html#765" class="Function">Congruent</a> <a id="2572" href="Function.Structures.Biased.html#680" class="Bound Operator">_≈₂_</a> <a id="2577" href="Function.Structures.Biased.html#620" class="Bound Operator">_≈₁_</a> <a id="2582" href="Function.Structures.Biased.html#2457" class="Bound">from</a>
|
||
<a id="IsStrictLeftInverse.strictlyInverseˡ"></a><a id="2591" href="Function.Structures.Biased.html#2591" class="Field">strictlyInverseˡ</a> <a id="2608" class="Symbol">:</a> <a id="2610" href="Function.Definitions.html#1622" class="Function">StrictlyInverseˡ</a> <a id="2627" href="Function.Structures.Biased.html#680" class="Bound Operator">_≈₂_</a> <a id="2632" href="Function.Structures.Biased.html#2444" class="Bound">to</a> <a id="2635" href="Function.Structures.Biased.html#2457" class="Bound">from</a>
|
||
|
||
<a id="IsStrictLeftInverse.isLeftInverse"></a><a id="2643" href="Function.Structures.Biased.html#2643" class="Function">isLeftInverse</a> <a id="2657" class="Symbol">:</a> <a id="2659" href="Function.Structures.html#2598" class="Record">IsLeftInverse</a> <a id="2673" href="Function.Structures.Biased.html#2444" class="Bound">to</a> <a id="2676" href="Function.Structures.Biased.html#2457" class="Bound">from</a>
|
||
<a id="2683" href="Function.Structures.Biased.html#2643" class="Function">isLeftInverse</a> <a id="2697" class="Symbol">=</a> <a id="2699" class="Keyword">record</a>
|
||
<a id="2710" class="Symbol">{</a> <a id="2712" href="Function.Structures.html#2682" class="Field">isCongruent</a> <a id="2724" class="Symbol">=</a> <a id="2726" href="Function.Structures.Biased.html#2513" class="Field">isCongruent</a>
|
||
<a id="2742" class="Symbol">;</a> <a id="2744" href="Function.Structures.html#2716" class="Field">from-cong</a> <a id="2754" class="Symbol">=</a> <a id="2756" href="Function.Structures.Biased.html#2547" class="Field">from-cong</a>
|
||
<a id="2770" class="Symbol">;</a> <a id="2772" href="Function.Structures.html#2760" class="Field">inverseˡ</a> <a id="2781" class="Symbol">=</a> <a id="2783" href="Function.Consequences.Setoid.html#2466" class="Function">strictlyInverseˡ⇒inverseˡ</a>
|
||
<a id="2817" href="Function.Structures.html#1207" class="Function">Eq₁.setoid</a> <a id="2828" href="Function.Structures.html#1351" class="Function">Eq₂.setoid</a> <a id="2839" href="Function.Structures.html#1062" class="Function">cong</a> <a id="2844" href="Function.Structures.Biased.html#2591" class="Field">strictlyInverseˡ</a>
|
||
<a id="2865" class="Symbol">}</a> <a id="2867" class="Keyword">where</a> <a id="2873" class="Keyword">open</a> <a id="2878" href="Function.Structures.html#995" class="Module">IsCongruent</a> <a id="2890" href="Function.Structures.Biased.html#2513" class="Field">isCongruent</a>
|
||
|
||
<a id="2903" class="Keyword">open</a> <a id="2908" href="Function.Structures.Biased.html#2423" class="Module">IsStrictLeftInverse</a> <a id="2928" class="Keyword">public</a>
|
||
<a id="2937" class="Keyword">using</a> <a id="2943" class="Symbol">()</a> <a id="2946" class="Keyword">renaming</a> <a id="2955" class="Symbol">(</a><a id="2956" href="Function.Structures.Biased.html#2643" class="Function">isLeftInverse</a> <a id="2970" class="Symbol">to</a> <a id="2973" class="Function">isStrictLeftInverse</a><a id="2992" class="Symbol">)</a>
|
||
|
||
<a id="2995" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="3068" class="Comment">-- Right inverse</a>
|
||
<a id="3085" class="Comment">------------------------------------------------------------------------</a>
|
||
|
||
<a id="3159" class="Keyword">record</a> <a id="IsStrictRightInverse"></a><a id="3166" href="Function.Structures.Biased.html#3166" class="Record">IsStrictRightInverse</a> <a id="3187" class="Symbol">(</a><a id="3188" href="Function.Structures.Biased.html#3188" class="Bound">to</a> <a id="3191" class="Symbol">:</a> <a id="3193" href="Function.Structures.Biased.html#608" class="Bound">A</a> <a id="3195" class="Symbol">→</a> <a id="3197" href="Function.Structures.Biased.html#668" class="Bound">B</a><a id="3198" class="Symbol">)</a> <a id="3200" class="Symbol">(</a><a id="3201" href="Function.Structures.Biased.html#3201" class="Bound">from</a> <a id="3206" class="Symbol">:</a> <a id="3208" href="Function.Structures.Biased.html#668" class="Bound">B</a> <a id="3210" class="Symbol">→</a> <a id="3212" href="Function.Structures.Biased.html#608" class="Bound">A</a><a id="3213" class="Symbol">)</a> <a id="3215" class="Symbol">:</a> <a id="3217" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="3221" class="Symbol">(</a><a id="3222" href="Function.Structures.Biased.html#594" class="Bound">a</a> <a id="3224" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="3226" href="Function.Structures.Biased.html#596" class="Bound">b</a> <a id="3228" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="3230" href="Function.Structures.Biased.html#598" class="Bound">ℓ₁</a> <a id="3233" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="3235" href="Function.Structures.Biased.html#601" class="Bound">ℓ₂</a><a id="3237" class="Symbol">)</a> <a id="3239" class="Keyword">where</a>
|
||
<a id="3247" class="Keyword">field</a>
|
||
<a id="IsStrictRightInverse.isCongruent"></a><a id="3257" href="Function.Structures.Biased.html#3257" class="Field">isCongruent</a> <a id="3269" class="Symbol">:</a> <a id="3271" href="Function.Structures.html#995" class="Record">IsCongruent</a> <a id="3283" href="Function.Structures.Biased.html#3188" class="Bound">to</a>
|
||
<a id="IsStrictRightInverse.from-cong"></a><a id="3290" href="Function.Structures.Biased.html#3290" class="Field">from-cong</a> <a id="3302" class="Symbol">:</a> <a id="3304" href="Function.Definitions.html#765" class="Function">Congruent</a> <a id="3314" href="Function.Structures.Biased.html#680" class="Bound Operator">_≈₂_</a> <a id="3319" href="Function.Structures.Biased.html#620" class="Bound Operator">_≈₁_</a> <a id="3324" href="Function.Structures.Biased.html#3201" class="Bound">from</a>
|
||
<a id="IsStrictRightInverse.strictlyInverseʳ"></a><a id="3333" href="Function.Structures.Biased.html#3333" class="Field">strictlyInverseʳ</a> <a id="3350" class="Symbol">:</a> <a id="3352" href="Function.Definitions.html#1726" class="Function">StrictlyInverseʳ</a> <a id="3369" href="Function.Structures.Biased.html#620" class="Bound Operator">_≈₁_</a> <a id="3374" href="Function.Structures.Biased.html#3188" class="Bound">to</a> <a id="3377" href="Function.Structures.Biased.html#3201" class="Bound">from</a>
|
||
|
||
<a id="IsStrictRightInverse.isRightInverse"></a><a id="3385" href="Function.Structures.Biased.html#3385" class="Function">isRightInverse</a> <a id="3400" class="Symbol">:</a> <a id="3402" href="Function.Structures.html#3113" class="Record">IsRightInverse</a> <a id="3417" href="Function.Structures.Biased.html#3188" class="Bound">to</a> <a id="3420" href="Function.Structures.Biased.html#3201" class="Bound">from</a>
|
||
<a id="3427" href="Function.Structures.Biased.html#3385" class="Function">isRightInverse</a> <a id="3442" class="Symbol">=</a> <a id="3444" class="Keyword">record</a>
|
||
<a id="3455" class="Symbol">{</a> <a id="3457" href="Function.Structures.html#3198" class="Field">isCongruent</a> <a id="3469" class="Symbol">=</a> <a id="3471" href="Function.Structures.Biased.html#3257" class="Field">isCongruent</a>
|
||
<a id="3487" class="Symbol">;</a> <a id="3489" href="Function.Structures.html#3231" class="Field">from-cong</a> <a id="3499" class="Symbol">=</a> <a id="3501" href="Function.Structures.Biased.html#3290" class="Field">from-cong</a>
|
||
<a id="3515" class="Symbol">;</a> <a id="3517" href="Function.Structures.html#3274" class="Field">inverseʳ</a> <a id="3526" class="Symbol">=</a> <a id="3528" href="Function.Consequences.Setoid.html#2953" class="Function">strictlyInverseʳ⇒inverseʳ</a>
|
||
<a id="3562" href="Function.Structures.html#1207" class="Function">Eq₁.setoid</a> <a id="3573" href="Function.Structures.html#1351" class="Function">Eq₂.setoid</a> <a id="3584" href="Function.Structures.Biased.html#3290" class="Field">from-cong</a> <a id="3594" href="Function.Structures.Biased.html#3333" class="Field">strictlyInverseʳ</a>
|
||
<a id="3615" class="Symbol">}</a> <a id="3617" class="Keyword">where</a> <a id="3623" class="Keyword">open</a> <a id="3628" href="Function.Structures.html#995" class="Module">IsCongruent</a> <a id="3640" href="Function.Structures.Biased.html#3257" class="Field">isCongruent</a>
|
||
|
||
<a id="3653" class="Keyword">open</a> <a id="3658" href="Function.Structures.Biased.html#3166" class="Module">IsStrictRightInverse</a> <a id="3679" class="Keyword">public</a>
|
||
<a id="3688" class="Keyword">using</a> <a id="3694" class="Symbol">()</a> <a id="3697" class="Keyword">renaming</a> <a id="3706" class="Symbol">(</a><a id="3707" href="Function.Structures.Biased.html#3385" class="Function">isRightInverse</a> <a id="3722" class="Symbol">to</a> <a id="3725" class="Function">isStrictRightInverse</a><a id="3745" class="Symbol">)</a>
|
||
|
||
<a id="3748" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="3821" class="Comment">-- Inverse</a>
|
||
<a id="3832" class="Comment">------------------------------------------------------------------------</a>
|
||
|
||
<a id="3906" class="Keyword">record</a> <a id="IsStrictInverse"></a><a id="3913" href="Function.Structures.Biased.html#3913" class="Record">IsStrictInverse</a> <a id="3929" class="Symbol">(</a><a id="3930" href="Function.Structures.Biased.html#3930" class="Bound">to</a> <a id="3933" class="Symbol">:</a> <a id="3935" href="Function.Structures.Biased.html#608" class="Bound">A</a> <a id="3937" class="Symbol">→</a> <a id="3939" href="Function.Structures.Biased.html#668" class="Bound">B</a><a id="3940" class="Symbol">)</a> <a id="3942" class="Symbol">(</a><a id="3943" href="Function.Structures.Biased.html#3943" class="Bound">from</a> <a id="3948" class="Symbol">:</a> <a id="3950" href="Function.Structures.Biased.html#668" class="Bound">B</a> <a id="3952" class="Symbol">→</a> <a id="3954" href="Function.Structures.Biased.html#608" class="Bound">A</a><a id="3955" class="Symbol">)</a> <a id="3957" class="Symbol">:</a> <a id="3959" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="3963" class="Symbol">(</a><a id="3964" href="Function.Structures.Biased.html#594" class="Bound">a</a> <a id="3966" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="3968" href="Function.Structures.Biased.html#596" class="Bound">b</a> <a id="3970" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="3972" href="Function.Structures.Biased.html#598" class="Bound">ℓ₁</a> <a id="3975" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="3977" href="Function.Structures.Biased.html#601" class="Bound">ℓ₂</a><a id="3979" class="Symbol">)</a> <a id="3981" class="Keyword">where</a>
|
||
<a id="3989" class="Keyword">field</a>
|
||
<a id="IsStrictInverse.isLeftInverse"></a><a id="3999" href="Function.Structures.Biased.html#3999" class="Field">isLeftInverse</a> <a id="4013" class="Symbol">:</a> <a id="4015" href="Function.Structures.html#2598" class="Record">IsLeftInverse</a> <a id="4029" href="Function.Structures.Biased.html#3930" class="Bound">to</a> <a id="4032" href="Function.Structures.Biased.html#3943" class="Bound">from</a>
|
||
<a id="IsStrictInverse.strictlyInverseʳ"></a><a id="4041" href="Function.Structures.Biased.html#4041" class="Field">strictlyInverseʳ</a> <a id="4058" class="Symbol">:</a> <a id="4060" href="Function.Definitions.html#1726" class="Function">StrictlyInverseʳ</a> <a id="4077" href="Function.Structures.Biased.html#620" class="Bound Operator">_≈₁_</a> <a id="4082" href="Function.Structures.Biased.html#3930" class="Bound">to</a> <a id="4085" href="Function.Structures.Biased.html#3943" class="Bound">from</a>
|
||
|
||
<a id="IsStrictInverse.isInverse"></a><a id="4093" href="Function.Structures.Biased.html#4093" class="Function">isInverse</a> <a id="4103" class="Symbol">:</a> <a id="4105" href="Function.Structures.html#3487" class="Record">IsInverse</a> <a id="4115" href="Function.Structures.Biased.html#3930" class="Bound">to</a> <a id="4118" href="Function.Structures.Biased.html#3943" class="Bound">from</a>
|
||
<a id="4125" href="Function.Structures.Biased.html#4093" class="Function">isInverse</a> <a id="4135" class="Symbol">=</a> <a id="4137" class="Keyword">record</a>
|
||
<a id="4148" class="Symbol">{</a> <a id="4150" href="Function.Structures.html#3567" class="Field">isLeftInverse</a> <a id="4164" class="Symbol">=</a> <a id="4166" href="Function.Structures.Biased.html#3999" class="Field">isLeftInverse</a>
|
||
<a id="4184" class="Symbol">;</a> <a id="4186" href="Function.Structures.html#3609" class="Field">inverseʳ</a> <a id="4200" class="Symbol">=</a> <a id="4202" href="Function.Consequences.Setoid.html#2953" class="Function">strictlyInverseʳ⇒inverseʳ</a>
|
||
<a id="4236" href="Function.Structures.html#1207" class="Function">Eq₁.setoid</a> <a id="4247" href="Function.Structures.html#1351" class="Function">Eq₂.setoid</a> <a id="4258" href="Function.Structures.html#2716" class="Function">from-cong</a> <a id="4268" href="Function.Structures.Biased.html#4041" class="Field">strictlyInverseʳ</a>
|
||
<a id="4289" class="Symbol">}</a> <a id="4291" class="Keyword">where</a> <a id="4297" class="Keyword">open</a> <a id="4302" href="Function.Structures.html#2598" class="Module">IsLeftInverse</a> <a id="4316" href="Function.Structures.Biased.html#3999" class="Field">isLeftInverse</a>
|
||
|
||
<a id="4331" class="Keyword">open</a> <a id="4336" href="Function.Structures.Biased.html#3913" class="Module">IsStrictInverse</a> <a id="4352" class="Keyword">public</a>
|
||
<a id="4361" class="Keyword">using</a> <a id="4367" class="Symbol">()</a> <a id="4370" class="Keyword">renaming</a> <a id="4379" class="Symbol">(</a><a id="4380" href="Function.Structures.Biased.html#4093" class="Function">isInverse</a> <a id="4390" class="Symbol">to</a> <a id="4393" class="Function">isStrictInverse</a><a id="4408" class="Symbol">)</a>
|
||
</pre></body></html> |