mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
239 lines
No EOL
112 KiB
HTML
239 lines
No EOL
112 KiB
HTML
<!DOCTYPE HTML>
|
||
<html><head><meta charset="utf-8"><title>Categories.Adjoint.Mate</title><link rel="stylesheet" href="Agda.css"></head><body><pre class="Agda"><a id="1" class="Symbol">{-#</a> <a id="5" class="Keyword">OPTIONS</a> <a id="13" class="Pragma">--without-K</a> <a id="25" class="Pragma">--safe</a> <a id="32" class="Symbol">#-}</a>
|
||
|
||
<a id="37" class="Comment">-- the usual notion of mate is defined by two isomorphisms between hom set(oid)s are natural,</a>
|
||
<a id="131" class="Comment">-- but due to explicit universe level, a different definition is used.</a>
|
||
<a id="202" class="Keyword">module</a> <a id="209" href="Categories.Adjoint.Mate.html" class="Module">Categories.Adjoint.Mate</a> <a id="233" class="Keyword">where</a>
|
||
|
||
<a id="240" class="Keyword">open</a> <a id="245" class="Keyword">import</a> <a id="252" href="Level.html" class="Module">Level</a>
|
||
<a id="258" class="Keyword">open</a> <a id="263" class="Keyword">import</a> <a id="270" href="Data.Product.html" class="Module">Data.Product</a> <a id="283" class="Keyword">using</a> <a id="289" class="Symbol">(</a><a id="290" href="Agda.Builtin.Sigma.html#165" class="Record">Σ</a><a id="291" class="Symbol">;</a> <a id="293" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">_,_</a><a id="296" class="Symbol">)</a>
|
||
<a id="298" class="Keyword">open</a> <a id="303" class="Keyword">import</a> <a id="310" href="Function.Bundles.html" class="Module">Function.Bundles</a> <a id="327" class="Keyword">using</a> <a id="333" class="Symbol">(</a><a id="334" href="Function.Bundles.html#2043" class="Record">Func</a><a id="338" class="Symbol">;</a> <a id="340" href="Function.Bundles.html#15133" class="Function Operator">_⟨$⟩_</a><a id="345" class="Symbol">)</a>
|
||
<a id="347" class="Keyword">open</a> <a id="352" class="Keyword">import</a> <a id="359" href="Function.Construct.Composition.html" class="Module">Function.Construct.Composition</a> <a id="390" class="Keyword">using</a> <a id="396" class="Symbol">(</a><a id="397" href="Function.Construct.Composition.html#5240" class="Function">function</a><a id="405" class="Symbol">)</a>
|
||
<a id="407" class="Keyword">open</a> <a id="412" class="Keyword">import</a> <a id="419" href="Function.Construct.Setoid.html" class="Module">Function.Construct.Setoid</a> <a id="445" class="Keyword">using</a> <a id="451" class="Symbol">()</a> <a id="454" class="Keyword">renaming</a> <a id="463" class="Symbol">(</a><a id="464" href="Function.Construct.Setoid.html#376" class="Function">setoid</a> <a id="471" class="Symbol">to</a> <a id="474" class="Function">_⇨_</a><a id="477" class="Symbol">)</a>
|
||
<a id="479" class="Keyword">open</a> <a id="484" class="Keyword">import</a> <a id="491" href="Relation.Binary.html" class="Module">Relation.Binary</a> <a id="507" class="Keyword">using</a> <a id="513" class="Symbol">(</a><a id="514" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a><a id="520" class="Symbol">;</a> <a id="522" href="Relation.Binary.Structures.html#1550" class="Record">IsEquivalence</a><a id="535" class="Symbol">)</a>
|
||
|
||
<a id="538" class="Keyword">open</a> <a id="543" class="Keyword">import</a> <a id="550" href="Categories.Category.html" class="Module">Categories.Category</a>
|
||
<a id="570" class="Keyword">open</a> <a id="575" class="Keyword">import</a> <a id="582" href="Categories.Category.Instance.Setoids.html" class="Module">Categories.Category.Instance.Setoids</a>
|
||
<a id="619" class="Keyword">open</a> <a id="624" class="Keyword">import</a> <a id="631" href="Categories.Functor.html" class="Module">Categories.Functor</a>
|
||
<a id="650" class="Keyword">open</a> <a id="655" class="Keyword">import</a> <a id="662" href="Categories.Functor.Hom.html" class="Module">Categories.Functor.Hom</a>
|
||
<a id="685" class="Keyword">open</a> <a id="690" class="Keyword">import</a> <a id="697" href="Categories.NaturalTransformation.html" class="Module">Categories.NaturalTransformation</a> <a id="730" class="Keyword">renaming</a> <a id="739" class="Symbol">(</a><a id="740" href="Categories.NaturalTransformation.Core.html#2132" class="Function">id</a> <a id="743" class="Symbol">to</a> <a id="746" class="Function">idN</a><a id="749" class="Symbol">)</a>
|
||
<a id="751" class="Keyword">open</a> <a id="756" class="Keyword">import</a> <a id="763" href="Categories.NaturalTransformation.Equivalence.html" class="Module">Categories.NaturalTransformation.Equivalence</a> <a id="808" class="Keyword">using</a> <a id="814" class="Symbol">(</a><a id="815" href="Categories.NaturalTransformation.Equivalence.html#630" class="Function Operator">_≃_</a><a id="818" class="Symbol">)</a>
|
||
<a id="820" class="Keyword">open</a> <a id="825" class="Keyword">import</a> <a id="832" href="Categories.Adjoint.html" class="Module">Categories.Adjoint</a>
|
||
<a id="851" class="Keyword">import</a> <a id="858" href="Categories.Morphism.Reasoning.html" class="Module">Categories.Morphism.Reasoning</a> <a id="888" class="Symbol">as</a> <a id="891" class="Module">MR</a>
|
||
|
||
<a id="895" class="Keyword">private</a>
|
||
<a id="905" class="Keyword">variable</a>
|
||
<a id="918" href="Categories.Adjoint.Mate.html#918" class="Generalizable">o</a> <a id="920" href="Categories.Adjoint.Mate.html#920" class="Generalizable">ℓ</a> <a id="922" href="Categories.Adjoint.Mate.html#922" class="Generalizable">e</a> <a id="924" class="Symbol">:</a> <a id="926" href="Agda.Primitive.html#742" class="Postulate">Level</a>
|
||
<a id="936" href="Categories.Adjoint.Mate.html#936" class="Generalizable">C</a> <a id="938" href="Categories.Adjoint.Mate.html#938" class="Generalizable">D</a> <a id="940" class="Symbol">:</a> <a id="942" href="Categories.Category.Core.html#442" class="Record">Category</a> <a id="951" href="Categories.Adjoint.Mate.html#918" class="Generalizable">o</a> <a id="953" href="Categories.Adjoint.Mate.html#920" class="Generalizable">ℓ</a> <a id="955" href="Categories.Adjoint.Mate.html#922" class="Generalizable">e</a>
|
||
<a id="961" href="Categories.Adjoint.Mate.html#961" class="Generalizable">L</a> <a id="963" href="Categories.Adjoint.Mate.html#963" class="Generalizable">L′</a> <a id="966" href="Categories.Adjoint.Mate.html#966" class="Generalizable">R</a> <a id="968" href="Categories.Adjoint.Mate.html#968" class="Generalizable">R′</a> <a id="971" class="Symbol">:</a> <a id="973" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="981" href="Categories.Adjoint.Mate.html#936" class="Generalizable">C</a> <a id="983" href="Categories.Adjoint.Mate.html#938" class="Generalizable">D</a>
|
||
|
||
|
||
<a id="987" class="Comment">-- this notion of mate can be seen in MacLane in which this notion is shown equivalent to the</a>
|
||
<a id="1081" class="Comment">-- definition via naturally isomorphic hom setoids.</a>
|
||
<a id="1133" class="Keyword">record</a> <a id="Mate"></a><a id="1140" href="Categories.Adjoint.Mate.html#1140" class="Record">Mate</a> <a id="1145" class="Symbol">{</a><a id="1146" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="1148" class="Symbol">:</a> <a id="1150" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="1158" href="Categories.Adjoint.Mate.html#936" class="Generalizable">C</a> <a id="1160" href="Categories.Adjoint.Mate.html#938" class="Generalizable">D</a><a id="1161" class="Symbol">}</a>
|
||
<a id="1175" class="Symbol">(</a><a id="1176" href="Categories.Adjoint.Mate.html#1176" class="Bound">L⊣R</a> <a id="1180" class="Symbol">:</a> <a id="1182" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="1184" href="Categories.Adjoint.html#7818" class="Function Operator">⊣</a> <a id="1186" href="Categories.Adjoint.Mate.html#966" class="Generalizable">R</a><a id="1187" class="Symbol">)</a> <a id="1189" class="Symbol">(</a><a id="1190" href="Categories.Adjoint.Mate.html#1190" class="Bound">L′⊣R′</a> <a id="1196" class="Symbol">:</a> <a id="1198" href="Categories.Adjoint.Mate.html#963" class="Generalizable">L′</a> <a id="1201" href="Categories.Adjoint.html#7818" class="Function Operator">⊣</a> <a id="1203" href="Categories.Adjoint.Mate.html#968" class="Generalizable">R′</a><a id="1205" class="Symbol">)</a>
|
||
<a id="1219" class="Symbol">(</a><a id="1220" href="Categories.Adjoint.Mate.html#1220" class="Bound">α</a> <a id="1222" class="Symbol">:</a> <a id="1224" href="Categories.NaturalTransformation.Core.html#466" class="Record">NaturalTransformation</a> <a id="1246" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="1248" href="Categories.Adjoint.Mate.html#963" class="Generalizable">L′</a><a id="1250" class="Symbol">)</a>
|
||
<a id="1264" class="Symbol">(</a><a id="1265" href="Categories.Adjoint.Mate.html#1265" class="Bound">β</a> <a id="1267" class="Symbol">:</a> <a id="1269" href="Categories.NaturalTransformation.Core.html#466" class="Record">NaturalTransformation</a> <a id="1291" href="Categories.Adjoint.Mate.html#968" class="Generalizable">R′</a> <a id="1294" href="Categories.Adjoint.Mate.html#966" class="Generalizable">R</a><a id="1295" class="Symbol">)</a> <a id="1297" class="Symbol">:</a> <a id="1299" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1303" class="Symbol">(</a><a id="1304" href="Level.html#602" class="Function">levelOfTerm</a> <a id="1316" href="Categories.Adjoint.Mate.html#1176" class="Bound">L⊣R</a> <a id="1320" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1322" href="Level.html#602" class="Function">levelOfTerm</a> <a id="1334" href="Categories.Adjoint.Mate.html#1190" class="Bound">L′⊣R′</a><a id="1339" class="Symbol">)</a> <a id="1341" class="Keyword">where</a>
|
||
<a id="1349" class="Keyword">private</a>
|
||
<a id="1361" class="Keyword">module</a> <a id="Mate.L⊣R"></a><a id="1368" href="Categories.Adjoint.Mate.html#1368" class="Module">L⊣R</a> <a id="1374" class="Symbol">=</a> <a id="1376" href="Categories.Adjoint.html#1260" class="Module">Adjoint</a> <a id="1384" href="Categories.Adjoint.Mate.html#1176" class="Bound">L⊣R</a>
|
||
<a id="1392" class="Keyword">module</a> <a id="Mate.L′⊣R′"></a><a id="1399" href="Categories.Adjoint.Mate.html#1399" class="Module">L′⊣R′</a> <a id="1405" class="Symbol">=</a> <a id="1407" href="Categories.Adjoint.html#1260" class="Module">Adjoint</a> <a id="1415" href="Categories.Adjoint.Mate.html#1190" class="Bound">L′⊣R′</a>
|
||
<a id="1425" class="Keyword">module</a> <a id="Mate.C"></a><a id="1432" href="Categories.Adjoint.Mate.html#1432" class="Module">C</a> <a id="1434" class="Symbol">=</a> <a id="1436" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="1445" href="Categories.Adjoint.Mate.html#1158" class="Bound">C</a>
|
||
<a id="1451" class="Keyword">module</a> <a id="Mate.D"></a><a id="1458" href="Categories.Adjoint.Mate.html#1458" class="Module">D</a> <a id="1460" class="Symbol">=</a> <a id="1462" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="1471" href="Categories.Adjoint.Mate.html#1160" class="Bound">D</a>
|
||
|
||
<a id="1476" class="Keyword">field</a>
|
||
<a id="Mate.commute₁"></a><a id="1486" href="Categories.Adjoint.Mate.html#1486" class="Field">commute₁</a> <a id="1495" class="Symbol">:</a> <a id="1497" class="Symbol">(</a><a id="1498" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a> <a id="1500" href="Categories.NaturalTransformation.Core.html#3439" class="Function Operator">∘ˡ</a> <a id="1503" href="Categories.Adjoint.Mate.html#1220" class="Bound">α</a><a id="1504" class="Symbol">)</a> <a id="1506" href="Categories.NaturalTransformation.Core.html#2439" class="Function Operator">∘ᵥ</a> <a id="1509" href="Categories.Adjoint.html#1473" class="Function">L⊣R.unit</a> <a id="1518" href="Categories.NaturalTransformation.Equivalence.html#630" class="Function Operator">≃</a> <a id="1520" class="Symbol">(</a><a id="1521" href="Categories.Adjoint.Mate.html#1265" class="Bound">β</a> <a id="1523" href="Categories.NaturalTransformation.Core.html#3784" class="Function Operator">∘ʳ</a> <a id="1526" href="Categories.Adjoint.Mate.html#1198" class="Bound">L′</a><a id="1528" class="Symbol">)</a> <a id="1530" href="Categories.NaturalTransformation.Core.html#2439" class="Function Operator">∘ᵥ</a> <a id="1533" href="Categories.Adjoint.html#1473" class="Function">L′⊣R′.unit</a>
|
||
<a id="Mate.commute₂"></a><a id="1548" href="Categories.Adjoint.Mate.html#1548" class="Field">commute₂</a> <a id="1557" class="Symbol">:</a> <a id="1559" href="Categories.Adjoint.html#1521" class="Function">L⊣R.counit</a> <a id="1570" href="Categories.NaturalTransformation.Core.html#2439" class="Function Operator">∘ᵥ</a> <a id="1573" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="1575" href="Categories.NaturalTransformation.Core.html#3439" class="Function Operator">∘ˡ</a> <a id="1578" href="Categories.Adjoint.Mate.html#1265" class="Bound">β</a> <a id="1580" href="Categories.NaturalTransformation.Equivalence.html#630" class="Function Operator">≃</a> <a id="1582" href="Categories.Adjoint.html#1521" class="Function">L′⊣R′.counit</a> <a id="1595" href="Categories.NaturalTransformation.Core.html#2439" class="Function Operator">∘ᵥ</a> <a id="1598" class="Symbol">(</a><a id="1599" href="Categories.Adjoint.Mate.html#1220" class="Bound">α</a> <a id="1601" href="Categories.NaturalTransformation.Core.html#3784" class="Function Operator">∘ʳ</a> <a id="1604" href="Categories.Adjoint.Mate.html#1203" class="Bound">R′</a><a id="1606" class="Symbol">)</a>
|
||
|
||
<a id="1611" class="Comment">-- there are two equivalent commutative diagram</a>
|
||
<a id="1661" class="Keyword">open</a> <a id="1666" href="Categories.NaturalTransformation.Core.html#466" class="Module">NaturalTransformation</a> <a id="1688" class="Keyword">renaming</a> <a id="1697" class="Symbol">(</a><a id="1698" href="Categories.NaturalTransformation.Core.html#827" class="Field">commute</a> <a id="1706" class="Symbol">to</a> <a id="1709" class="Field">η-commute</a><a id="1718" class="Symbol">)</a>
|
||
<a id="1722" class="Keyword">open</a> <a id="1727" href="Categories.Functor.Core.html#248" class="Module">Functor</a>
|
||
<a id="1737" class="Keyword">module</a> <a id="1744" href="Categories.Adjoint.Mate.html#1744" class="Module">_</a> <a id="1746" class="Keyword">where</a>
|
||
<a id="1756" class="Keyword">open</a> <a id="1761" href="Categories.Adjoint.Mate.html#1458" class="Module">D</a>
|
||
<a id="1767" class="Keyword">open</a> <a id="1772" href="Categories.Category.Core.html#2462" class="Module">HomReasoning</a>
|
||
<a id="1789" class="Keyword">open</a> <a id="1794" href="Categories.Morphism.Reasoning.html" class="Module">MR</a> <a id="1797" href="Categories.Adjoint.Mate.html#1160" class="Bound">D</a>
|
||
|
||
<a id="1804" href="Categories.Adjoint.Mate.html#1804" class="Function">commute₃</a> <a id="1813" class="Symbol">:</a> <a id="1815" class="Symbol">∀</a> <a id="1817" class="Symbol">{</a><a id="1818" href="Categories.Adjoint.Mate.html#1818" class="Bound">X</a><a id="1819" class="Symbol">}</a> <a id="1821" class="Symbol">→</a> <a id="1823" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="1836" class="Symbol">(</a><a id="1837" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="1840" href="Categories.Adjoint.Mate.html#1198" class="Bound">L′</a> <a id="1843" href="Categories.Adjoint.Mate.html#1818" class="Bound">X</a><a id="1844" class="Symbol">)</a> <a id="1846" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="1848" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1851" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="1853" class="Symbol">(</a><a id="1854" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="1856" href="Categories.Adjoint.Mate.html#1265" class="Bound">β</a> <a id="1858" class="Symbol">(</a><a id="1859" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="1862" href="Categories.Adjoint.Mate.html#1198" class="Bound">L′</a> <a id="1865" href="Categories.Adjoint.Mate.html#1818" class="Bound">X</a><a id="1866" class="Symbol">))</a> <a id="1869" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="1871" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1874" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="1876" class="Symbol">(</a><a id="1877" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="1890" href="Categories.Adjoint.Mate.html#1818" class="Bound">X</a><a id="1891" class="Symbol">)</a> <a id="1893" href="Categories.Category.Core.html#595" class="Function Operator">≈</a> <a id="1895" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="1897" href="Categories.Adjoint.Mate.html#1220" class="Bound">α</a> <a id="1899" href="Categories.Adjoint.Mate.html#1818" class="Bound">X</a>
|
||
<a id="1905" href="Categories.Adjoint.Mate.html#1804" class="Function">commute₃</a> <a id="1914" class="Symbol">{</a><a id="1915" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="1916" class="Symbol">}</a> <a id="1918" class="Symbol">=</a> <a id="1920" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="1932" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="1945" class="Symbol">(</a><a id="1946" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="1949" href="Categories.Adjoint.Mate.html#1198" class="Bound">L′</a> <a id="1952" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="1953" class="Symbol">)</a> <a id="1955" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="1957" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1960" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="1962" class="Symbol">(</a><a id="1963" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="1965" href="Categories.Adjoint.Mate.html#1265" class="Bound">β</a> <a id="1967" class="Symbol">(</a><a id="1968" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="1971" href="Categories.Adjoint.Mate.html#1198" class="Bound">L′</a> <a id="1974" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="1975" class="Symbol">))</a> <a id="1978" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="1980" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1983" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="1985" class="Symbol">(</a><a id="1986" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="1999" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="2000" class="Symbol">)</a>
|
||
<a id="2010" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="2014" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="2022" href="Categories.Functor.Core.html#565" class="Field">homomorphism</a> <a id="2035" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="2037" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="2045" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="2058" class="Symbol">(</a><a id="2059" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2062" href="Categories.Adjoint.Mate.html#1198" class="Bound">L′</a> <a id="2065" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="2066" class="Symbol">)</a> <a id="2068" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="2070" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2073" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="2075" class="Symbol">(</a><a id="2076" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="2078" href="Categories.Adjoint.Mate.html#1265" class="Bound">β</a> <a id="2080" class="Symbol">(</a><a id="2081" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2084" href="Categories.Adjoint.Mate.html#1198" class="Bound">L′</a> <a id="2087" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="2088" class="Symbol">)</a> <a id="2090" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="2094" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="2107" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="2108" class="Symbol">)</a>
|
||
<a id="2118" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="2122" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="2130" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="2139" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="2141" href="Categories.Adjoint.Mate.html#1486" class="Field">commute₁</a> <a id="2150" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="2158" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="2171" class="Symbol">(</a><a id="2172" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2175" href="Categories.Adjoint.Mate.html#1198" class="Bound">L′</a> <a id="2178" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="2179" class="Symbol">)</a> <a id="2181" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="2183" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2186" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="2188" class="Symbol">(</a><a id="2189" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2192" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a> <a id="2194" class="Symbol">(</a><a id="2195" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="2197" href="Categories.Adjoint.Mate.html#1220" class="Bound">α</a> <a id="2199" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="2200" class="Symbol">)</a> <a id="2202" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="2206" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="2217" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a><a id="2218" class="Symbol">)</a>
|
||
<a id="2228" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2231" href="Categories.Adjoint.html#2026" class="Function">L⊣R.RLadjunct≈id</a> <a id="2248" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2256" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="2258" href="Categories.Adjoint.Mate.html#1220" class="Bound">α</a> <a id="2260" href="Categories.Adjoint.Mate.html#1915" class="Bound">X</a>
|
||
<a id="2270" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="2275" class="Keyword">module</a> <a id="2282" href="Categories.Adjoint.Mate.html#2282" class="Module">_</a> <a id="2284" class="Keyword">where</a>
|
||
<a id="2294" class="Keyword">open</a> <a id="2299" href="Categories.Adjoint.Mate.html#1432" class="Module">C</a>
|
||
<a id="2305" class="Keyword">open</a> <a id="2310" href="Categories.Category.Core.html#2462" class="Module">HomReasoning</a>
|
||
<a id="2327" class="Keyword">open</a> <a id="2332" href="Categories.Morphism.Reasoning.html" class="Module">MR</a> <a id="2335" href="Categories.Adjoint.Mate.html#1158" class="Bound">C</a>
|
||
|
||
<a id="2342" href="Categories.Adjoint.Mate.html#2342" class="Function">commute₄</a> <a id="2351" class="Symbol">:</a> <a id="2353" class="Symbol">∀</a> <a id="2355" class="Symbol">{</a><a id="2356" href="Categories.Adjoint.Mate.html#2356" class="Bound">X</a><a id="2357" class="Symbol">}</a> <a id="2359" class="Symbol">→</a> <a id="2361" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2364" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a> <a id="2366" class="Symbol">(</a><a id="2367" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="2382" href="Categories.Adjoint.Mate.html#2356" class="Bound">X</a><a id="2383" class="Symbol">)</a> <a id="2385" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="2387" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2390" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a> <a id="2392" class="Symbol">(</a><a id="2393" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="2395" href="Categories.Adjoint.Mate.html#1220" class="Bound">α</a> <a id="2397" class="Symbol">(</a><a id="2398" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2401" href="Categories.Adjoint.Mate.html#1203" class="Bound">R′</a> <a id="2404" href="Categories.Adjoint.Mate.html#2356" class="Bound">X</a><a id="2405" class="Symbol">))</a> <a id="2408" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="2410" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="2421" class="Symbol">(</a><a id="2422" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2425" href="Categories.Adjoint.Mate.html#1203" class="Bound">R′</a> <a id="2428" href="Categories.Adjoint.Mate.html#2356" class="Bound">X</a><a id="2429" class="Symbol">)</a> <a id="2431" href="Categories.Category.Core.html#595" class="Function Operator">≈</a> <a id="2433" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="2435" href="Categories.Adjoint.Mate.html#1265" class="Bound">β</a> <a id="2437" href="Categories.Adjoint.Mate.html#2356" class="Bound">X</a>
|
||
<a id="2443" href="Categories.Adjoint.Mate.html#2342" class="Function">commute₄</a> <a id="2452" class="Symbol">{</a><a id="2453" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a><a id="2454" class="Symbol">}</a> <a id="2456" class="Symbol">=</a> <a id="2458" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="2470" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2473" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a> <a id="2475" class="Symbol">(</a><a id="2476" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="2491" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a><a id="2492" class="Symbol">)</a> <a id="2494" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="2496" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2499" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a> <a id="2501" class="Symbol">(</a><a id="2502" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="2504" href="Categories.Adjoint.Mate.html#1220" class="Bound">α</a> <a id="2506" class="Symbol">(</a><a id="2507" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2510" href="Categories.Adjoint.Mate.html#1203" class="Bound">R′</a> <a id="2513" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a><a id="2514" class="Symbol">))</a> <a id="2517" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="2519" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="2530" class="Symbol">(</a><a id="2531" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2534" href="Categories.Adjoint.Mate.html#1203" class="Bound">R′</a> <a id="2537" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a><a id="2538" class="Symbol">)</a>
|
||
<a id="2548" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="2552" href="Categories.Morphism.Reasoning.Core.html#2485" class="Function">pushˡ</a> <a id="2558" class="Symbol">(</a><a id="2559" href="Categories.Functor.Core.html#565" class="Field">homomorphism</a> <a id="2572" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a><a id="2573" class="Symbol">)</a> <a id="2575" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="2583" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2586" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a> <a id="2588" class="Symbol">(</a><a id="2589" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="2604" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a> <a id="2606" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="2610" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="2612" href="Categories.Adjoint.Mate.html#1220" class="Bound">α</a> <a id="2614" class="Symbol">(</a><a id="2615" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2618" href="Categories.Adjoint.Mate.html#1203" class="Bound">R′</a> <a id="2621" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a><a id="2622" class="Symbol">))</a> <a id="2625" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="2627" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="2638" class="Symbol">(</a><a id="2639" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2642" href="Categories.Adjoint.Mate.html#1203" class="Bound">R′</a> <a id="2645" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a><a id="2646" class="Symbol">)</a>
|
||
<a id="2656" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="2660" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="2669" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a> <a id="2671" href="Categories.Adjoint.Mate.html#1548" class="Field">commute₂</a> <a id="2680" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="2688" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="2696" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2699" href="Categories.Adjoint.Mate.html#1186" class="Bound">R</a> <a id="2701" class="Symbol">(</a><a id="2702" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="2715" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a> <a id="2717" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="2721" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="2724" href="Categories.Adjoint.Mate.html#1146" class="Bound">L</a> <a id="2726" class="Symbol">(</a><a id="2727" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="2729" href="Categories.Adjoint.Mate.html#1265" class="Bound">β</a> <a id="2731" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a><a id="2732" class="Symbol">))</a> <a id="2735" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="2737" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="2748" class="Symbol">(</a><a id="2749" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="2752" href="Categories.Adjoint.Mate.html#1203" class="Bound">R′</a> <a id="2755" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a><a id="2756" class="Symbol">)</a>
|
||
<a id="2766" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2769" href="Categories.Adjoint.html#2533" class="Function">L⊣R.LRadjunct≈id</a> <a id="2786" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2794" href="Categories.NaturalTransformation.Core.html#783" class="Field">η</a> <a id="2796" href="Categories.Adjoint.Mate.html#1265" class="Bound">β</a> <a id="2798" href="Categories.Adjoint.Mate.html#2453" class="Bound">X</a>
|
||
<a id="2809" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="2812" class="Keyword">record</a> <a id="HaveMate"></a><a id="2819" href="Categories.Adjoint.Mate.html#2819" class="Record">HaveMate</a> <a id="2828" class="Symbol">{</a><a id="2829" href="Categories.Adjoint.Mate.html#2829" class="Bound">L</a> <a id="2831" href="Categories.Adjoint.Mate.html#2831" class="Bound">L′</a> <a id="2834" class="Symbol">:</a> <a id="2836" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="2844" href="Categories.Adjoint.Mate.html#936" class="Generalizable">C</a> <a id="2846" href="Categories.Adjoint.Mate.html#938" class="Generalizable">D</a><a id="2847" class="Symbol">}</a> <a id="2849" class="Symbol">{</a><a id="2850" href="Categories.Adjoint.Mate.html#2850" class="Bound">R</a> <a id="2852" href="Categories.Adjoint.Mate.html#2852" class="Bound">R′</a> <a id="2855" class="Symbol">:</a> <a id="2857" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="2865" href="Categories.Adjoint.Mate.html#938" class="Generalizable">D</a> <a id="2867" href="Categories.Adjoint.Mate.html#936" class="Generalizable">C</a><a id="2868" class="Symbol">}</a>
|
||
<a id="2886" class="Symbol">(</a><a id="2887" href="Categories.Adjoint.Mate.html#2887" class="Bound">L⊣R</a> <a id="2891" class="Symbol">:</a> <a id="2893" href="Categories.Adjoint.Mate.html#2829" class="Bound">L</a> <a id="2895" href="Categories.Adjoint.html#7818" class="Function Operator">⊣</a> <a id="2897" href="Categories.Adjoint.Mate.html#2850" class="Bound">R</a><a id="2898" class="Symbol">)</a> <a id="2900" class="Symbol">(</a><a id="2901" href="Categories.Adjoint.Mate.html#2901" class="Bound">L′⊣R′</a> <a id="2907" class="Symbol">:</a> <a id="2909" href="Categories.Adjoint.Mate.html#2831" class="Bound">L′</a> <a id="2912" href="Categories.Adjoint.html#7818" class="Function Operator">⊣</a> <a id="2914" href="Categories.Adjoint.Mate.html#2852" class="Bound">R′</a><a id="2916" class="Symbol">)</a> <a id="2918" class="Symbol">:</a> <a id="2920" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="2924" class="Symbol">(</a><a id="2925" href="Level.html#602" class="Function">levelOfTerm</a> <a id="2937" href="Categories.Adjoint.Mate.html#2887" class="Bound">L⊣R</a> <a id="2941" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2943" href="Level.html#602" class="Function">levelOfTerm</a> <a id="2955" href="Categories.Adjoint.Mate.html#2901" class="Bound">L′⊣R′</a><a id="2960" class="Symbol">)</a> <a id="2962" class="Keyword">where</a>
|
||
<a id="2970" class="Keyword">field</a>
|
||
<a id="HaveMate.α"></a><a id="2980" href="Categories.Adjoint.Mate.html#2980" class="Field">α</a> <a id="2985" class="Symbol">:</a> <a id="2987" href="Categories.NaturalTransformation.Core.html#466" class="Record">NaturalTransformation</a> <a id="3009" href="Categories.Adjoint.Mate.html#2829" class="Bound">L</a> <a id="3011" href="Categories.Adjoint.Mate.html#2831" class="Bound">L′</a>
|
||
<a id="HaveMate.β"></a><a id="3018" href="Categories.Adjoint.Mate.html#3018" class="Field">β</a> <a id="3023" class="Symbol">:</a> <a id="3025" href="Categories.NaturalTransformation.Core.html#466" class="Record">NaturalTransformation</a> <a id="3047" href="Categories.Adjoint.Mate.html#2852" class="Bound">R′</a> <a id="3050" href="Categories.Adjoint.Mate.html#2850" class="Bound">R</a>
|
||
<a id="HaveMate.mate"></a><a id="3056" href="Categories.Adjoint.Mate.html#3056" class="Field">mate</a> <a id="3061" class="Symbol">:</a> <a id="3063" href="Categories.Adjoint.Mate.html#1140" class="Record">Mate</a> <a id="3068" href="Categories.Adjoint.Mate.html#2887" class="Bound">L⊣R</a> <a id="3072" href="Categories.Adjoint.Mate.html#2901" class="Bound">L′⊣R′</a> <a id="3078" href="Categories.Adjoint.Mate.html#2980" class="Field">α</a> <a id="3080" href="Categories.Adjoint.Mate.html#3018" class="Field">β</a>
|
||
|
||
<a id="3085" class="Keyword">module</a> <a id="HaveMate.α"></a><a id="3092" href="Categories.Adjoint.Mate.html#3092" class="Module">α</a> <a id="3094" class="Symbol">=</a> <a id="3096" href="Categories.NaturalTransformation.Core.html#466" class="Module">NaturalTransformation</a> <a id="3118" href="Categories.Adjoint.Mate.html#2980" class="Field">α</a>
|
||
<a id="3122" class="Keyword">module</a> <a id="HaveMate.β"></a><a id="3129" href="Categories.Adjoint.Mate.html#3129" class="Module">β</a> <a id="3131" class="Symbol">=</a> <a id="3133" href="Categories.NaturalTransformation.Core.html#466" class="Module">NaturalTransformation</a> <a id="3155" href="Categories.Adjoint.Mate.html#3018" class="Field">β</a>
|
||
<a id="3159" class="Keyword">open</a> <a id="3164" href="Categories.Adjoint.Mate.html#1140" class="Module">Mate</a> <a id="3169" href="Categories.Adjoint.Mate.html#3056" class="Field">mate</a> <a id="3174" class="Keyword">public</a>
|
||
|
||
<a id="3182" class="Comment">-- show that the commutative diagram implies natural isomorphism between homsetoids.</a>
|
||
<a id="3267" class="Comment">-- the problem is that two homsetoids live in two universe level, in a situation similar to the definition</a>
|
||
<a id="3374" class="Comment">-- of adjoint via naturally isomorphic homsetoids.</a>
|
||
<a id="3425" class="Keyword">module</a> <a id="3432" href="Categories.Adjoint.Mate.html#3432" class="Module">_</a> <a id="3434" class="Symbol">{</a><a id="3435" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a> <a id="3437" href="Categories.Adjoint.Mate.html#3437" class="Bound">L′</a> <a id="3440" class="Symbol">:</a> <a id="3442" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="3450" href="Categories.Adjoint.Mate.html#936" class="Generalizable">C</a> <a id="3452" href="Categories.Adjoint.Mate.html#938" class="Generalizable">D</a><a id="3453" class="Symbol">}</a> <a id="3455" class="Symbol">{</a><a id="3456" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a> <a id="3458" href="Categories.Adjoint.Mate.html#3458" class="Bound">R′</a> <a id="3461" class="Symbol">:</a> <a id="3463" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="3471" href="Categories.Adjoint.Mate.html#938" class="Generalizable">D</a> <a id="3473" href="Categories.Adjoint.Mate.html#936" class="Generalizable">C</a><a id="3474" class="Symbol">}</a>
|
||
<a id="3485" class="Symbol">{</a><a id="3486" href="Categories.Adjoint.Mate.html#3486" class="Bound">L⊣R</a> <a id="3490" class="Symbol">:</a> <a id="3492" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a> <a id="3494" href="Categories.Adjoint.html#7818" class="Function Operator">⊣</a> <a id="3496" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a><a id="3497" class="Symbol">}</a> <a id="3499" class="Symbol">{</a><a id="3500" href="Categories.Adjoint.Mate.html#3500" class="Bound">L′⊣R′</a> <a id="3506" class="Symbol">:</a> <a id="3508" href="Categories.Adjoint.Mate.html#3437" class="Bound">L′</a> <a id="3511" href="Categories.Adjoint.html#7818" class="Function Operator">⊣</a> <a id="3513" href="Categories.Adjoint.Mate.html#3458" class="Bound">R′</a><a id="3515" class="Symbol">}</a>
|
||
<a id="3526" class="Symbol">{</a><a id="3527" href="Categories.Adjoint.Mate.html#3527" class="Bound">α</a> <a id="3529" class="Symbol">:</a> <a id="3531" href="Categories.NaturalTransformation.Core.html#466" class="Record">NaturalTransformation</a> <a id="3553" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a> <a id="3555" href="Categories.Adjoint.Mate.html#3437" class="Bound">L′</a><a id="3557" class="Symbol">}</a>
|
||
<a id="3568" class="Symbol">{</a><a id="3569" href="Categories.Adjoint.Mate.html#3569" class="Bound">β</a> <a id="3571" class="Symbol">:</a> <a id="3573" href="Categories.NaturalTransformation.Core.html#466" class="Record">NaturalTransformation</a> <a id="3595" href="Categories.Adjoint.Mate.html#3458" class="Bound">R′</a> <a id="3598" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a><a id="3599" class="Symbol">}</a>
|
||
<a id="3610" class="Symbol">(</a><a id="3611" href="Categories.Adjoint.Mate.html#3611" class="Bound">mate</a> <a id="3616" class="Symbol">:</a> <a id="3618" href="Categories.Adjoint.Mate.html#1140" class="Record">Mate</a> <a id="3623" href="Categories.Adjoint.Mate.html#3486" class="Bound">L⊣R</a> <a id="3627" href="Categories.Adjoint.Mate.html#3500" class="Bound">L′⊣R′</a> <a id="3633" href="Categories.Adjoint.Mate.html#3527" class="Bound">α</a> <a id="3635" href="Categories.Adjoint.Mate.html#3569" class="Bound">β</a><a id="3636" class="Symbol">)</a> <a id="3638" class="Keyword">where</a>
|
||
<a id="3646" class="Keyword">private</a>
|
||
<a id="3658" class="Keyword">open</a> <a id="3663" href="Categories.Adjoint.Mate.html#1140" class="Module">Mate</a> <a id="3668" href="Categories.Adjoint.Mate.html#3611" class="Bound">mate</a>
|
||
<a id="3677" class="Keyword">open</a> <a id="3682" href="Categories.Functor.Core.html#248" class="Module">Functor</a>
|
||
<a id="3694" class="Keyword">module</a> <a id="3701" href="Categories.Adjoint.Mate.html#3701" class="Module">C</a> <a id="3707" class="Symbol">=</a> <a id="3709" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="3718" href="Categories.Adjoint.Mate.html#3450" class="Bound">C</a>
|
||
<a id="3724" class="Keyword">module</a> <a id="3731" href="Categories.Adjoint.Mate.html#3731" class="Module">D</a> <a id="3737" class="Symbol">=</a> <a id="3739" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="3748" href="Categories.Adjoint.Mate.html#3452" class="Bound">D</a>
|
||
<a id="3754" class="Keyword">module</a> <a id="3761" href="Categories.Adjoint.Mate.html#3761" class="Module">α</a> <a id="3767" class="Symbol">=</a> <a id="3769" href="Categories.NaturalTransformation.Core.html#466" class="Module">NaturalTransformation</a> <a id="3791" href="Categories.Adjoint.Mate.html#3527" class="Bound">α</a>
|
||
<a id="3797" class="Keyword">module</a> <a id="3804" href="Categories.Adjoint.Mate.html#3804" class="Module">β</a> <a id="3810" class="Symbol">=</a> <a id="3812" href="Categories.NaturalTransformation.Core.html#466" class="Module">NaturalTransformation</a> <a id="3834" href="Categories.Adjoint.Mate.html#3569" class="Bound">β</a>
|
||
<a id="3840" class="Keyword">module</a> <a id="3847" href="Categories.Adjoint.Mate.html#3847" class="Module">L⊣R</a> <a id="3853" class="Symbol">=</a> <a id="3855" href="Categories.Adjoint.html#1260" class="Module">Adjoint</a> <a id="3863" href="Categories.Adjoint.Mate.html#3486" class="Bound">L⊣R</a>
|
||
<a id="3871" class="Keyword">module</a> <a id="3878" href="Categories.Adjoint.Mate.html#3878" class="Module">L′⊣R′</a> <a id="3884" class="Symbol">=</a> <a id="3886" href="Categories.Adjoint.html#1260" class="Module">Adjoint</a> <a id="3894" href="Categories.Adjoint.Mate.html#3500" class="Bound">L′⊣R′</a>
|
||
|
||
<a id="3903" class="Comment">-- there are two squares to show</a>
|
||
<a id="3938" class="Keyword">module</a> <a id="3945" href="Categories.Adjoint.Mate.html#3945" class="Module">_</a> <a id="3947" class="Symbol">{</a><a id="3948" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a> <a id="3950" class="Symbol">:</a> <a id="3952" href="Categories.Category.Core.html#559" class="Function">C.Obj</a><a id="3957" class="Symbol">}</a> <a id="3959" class="Symbol">{</a><a id="3960" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a> <a id="3962" class="Symbol">:</a> <a id="3964" href="Categories.Category.Core.html#559" class="Function">D.Obj</a><a id="3969" class="Symbol">}</a> <a id="3971" class="Keyword">where</a>
|
||
<a id="3981" class="Keyword">private</a>
|
||
<a id="3995" href="Categories.Adjoint.Mate.html#3995" class="Function">From</a> <a id="4000" class="Symbol">:</a> <a id="4002" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="4009" class="Symbol">_</a> <a id="4011" class="Symbol">_</a>
|
||
<a id="4019" href="Categories.Adjoint.Mate.html#3995" class="Function">From</a> <a id="4024" class="Symbol">=</a> <a id="4026" href="Categories.Functor.Core.html#432" class="Function">L′⊣R′.Hom[L-,-].F₀</a> <a id="4045" class="Symbol">(</a><a id="4046" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a> <a id="4048" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="4050" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a><a id="4051" class="Symbol">)</a>
|
||
<a id="4059" href="Categories.Adjoint.Mate.html#4059" class="Function">To</a> <a id="4062" class="Symbol">:</a> <a id="4064" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="4071" class="Symbol">_</a> <a id="4073" class="Symbol">_</a>
|
||
<a id="4081" href="Categories.Adjoint.Mate.html#4059" class="Function">To</a> <a id="4084" class="Symbol">=</a> <a id="4086" href="Categories.Functor.Core.html#432" class="Function">L⊣R.Hom[-,R-].F₀</a> <a id="4103" class="Symbol">(</a><a id="4104" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a> <a id="4106" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="4108" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a><a id="4109" class="Symbol">)</a>
|
||
<a id="4117" href="Categories.Adjoint.Mate.html#4117" class="Function">SS</a> <a id="4120" class="Symbol">=</a> <a id="4122" href="Categories.Adjoint.Mate.html#3995" class="Function">From</a> <a id="4127" href="Categories.Adjoint.Mate.html#474" class="Function Operator">⇨</a> <a id="4129" href="Categories.Adjoint.Mate.html#4059" class="Function">To</a>
|
||
<a id="4136" class="Keyword">open</a> <a id="4141" href="Relation.Binary.Bundles.html#1080" class="Module">Setoid</a> <a id="4148" href="Categories.Adjoint.Mate.html#4117" class="Function">SS</a> <a id="4151" class="Keyword">using</a> <a id="4157" class="Symbol">(</a><a id="4158" href="Relation.Binary.Bundles.html#1169" class="Field Operator">_≈_</a><a id="4161" class="Symbol">)</a>
|
||
<a id="4167" class="Keyword">open</a> <a id="4172" href="Categories.Adjoint.Mate.html#3701" class="Module">C</a> <a id="4174" class="Keyword">hiding</a> <a id="4181" class="Symbol">(</a><a id="4182" href="Categories.Category.Core.html#595" class="Function Operator">_≈_</a><a id="4185" class="Symbol">)</a>
|
||
<a id="4191" class="Keyword">open</a> <a id="4196" href="Categories.Morphism.Reasoning.html" class="Module">MR</a> <a id="4199" href="Categories.Adjoint.Mate.html#3450" class="Bound">C</a>
|
||
<a id="4205" class="Keyword">open</a> <a id="4210" href="Categories.Category.Core.html#2462" class="Module">C.HomReasoning</a>
|
||
<a id="4229" class="Keyword">module</a> <a id="4236" href="Categories.Adjoint.Mate.html#4236" class="Module">DH</a> <a id="4239" class="Symbol">=</a> <a id="4241" href="Categories.Category.Core.html#2462" class="Module">D.HomReasoning</a>
|
||
<a id="4260" class="Keyword">private</a>
|
||
<a id="4274" class="Comment">-- annoyingly the new bundles don't export this</a>
|
||
<a id="4328" href="Categories.Adjoint.Mate.html#4328" class="Function">D⟶C</a> <a id="4332" class="Symbol">:</a> <a id="4334" href="Function.Bundles.html#2043" class="Record">Func</a> <a id="4339" class="Symbol">(</a><a id="4340" href="Categories.Category.Core.html#1815" class="Function">D.hom-setoid</a> <a id="4353" class="Symbol">{</a><a id="4354" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="4357" href="Categories.Adjoint.Mate.html#3437" class="Bound">L′</a> <a id="4360" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a><a id="4361" class="Symbol">}</a> <a id="4363" class="Symbol">{</a><a id="4364" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a><a id="4365" class="Symbol">})</a> <a id="4368" class="Symbol">(</a><a id="4369" href="Categories.Category.Core.html#1815" class="Function">C.hom-setoid</a> <a id="4382" class="Symbol">{</a><a id="4383" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a><a id="4384" class="Symbol">}</a> <a id="4386" class="Symbol">{</a><a id="4387" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="4390" href="Categories.Adjoint.Mate.html#3458" class="Bound">R′</a> <a id="4393" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a><a id="4394" class="Symbol">})</a>
|
||
<a id="4403" href="Categories.Adjoint.Mate.html#4328" class="Function">D⟶C</a> <a id="4407" class="Symbol">=</a> <a id="4409" class="Keyword">record</a>
|
||
<a id="4424" class="Symbol">{</a> <a id="4426" href="Function.Bundles.html#2094" class="Field">to</a> <a id="4429" class="Symbol">=</a> <a id="4431" href="Function.Bundles.html#7394" class="Function">L′⊣R′.Hom-inverse.to</a> <a id="4452" class="Symbol">{</a><a id="4453" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a><a id="4454" class="Symbol">}</a> <a id="4456" class="Symbol">{</a><a id="4457" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a><a id="4458" class="Symbol">}</a>
|
||
<a id="4468" class="Symbol">;</a> <a id="4470" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="4475" class="Symbol">=</a> <a id="4477" href="Function.Bundles.html#7442" class="Function">L′⊣R′.Hom-inverse.to-cong</a>
|
||
<a id="4511" class="Symbol">}</a>
|
||
<a id="4519" href="Categories.Adjoint.Mate.html#4519" class="Function">D⟶C′</a> <a id="4524" class="Symbol">:</a> <a id="4526" href="Function.Bundles.html#2043" class="Record">Func</a> <a id="4531" href="Categories.Category.Core.html#1815" class="Function">D.hom-setoid</a> <a id="4544" href="Categories.Category.Core.html#1815" class="Function">C.hom-setoid</a>
|
||
<a id="4563" href="Categories.Adjoint.Mate.html#4519" class="Function">D⟶C′</a> <a id="4568" class="Symbol">=</a> <a id="4570" class="Keyword">record</a>
|
||
<a id="4585" class="Symbol">{</a> <a id="4587" href="Function.Bundles.html#2094" class="Field">to</a> <a id="4590" class="Symbol">=</a> <a id="4592" href="Function.Bundles.html#7394" class="Function">L⊣R.Hom-inverse.to</a> <a id="4611" class="Symbol">{</a><a id="4612" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a><a id="4613" class="Symbol">}</a> <a id="4615" class="Symbol">{</a><a id="4616" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a><a id="4617" class="Symbol">}</a>
|
||
<a id="4627" class="Symbol">;</a> <a id="4629" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="4634" class="Symbol">=</a> <a id="4636" href="Function.Bundles.html#7442" class="Function">L⊣R.Hom-inverse.to-cong</a>
|
||
<a id="4668" class="Symbol">}</a>
|
||
|
||
<a id="4675" href="Categories.Adjoint.Mate.html#4675" class="Function">mate-commute₁</a> <a id="4689" class="Symbol">:</a> <a id="4691" href="Function.Construct.Composition.html#5240" class="Function">function</a> <a id="4700" href="Categories.Adjoint.Mate.html#4328" class="Function">D⟶C</a> <a id="4704" class="Symbol">(</a><a id="4705" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="4708" href="Categories.Functor.Hom.html#1578" class="Function Operator">Hom[</a> <a id="4713" href="Categories.Adjoint.Mate.html#3450" class="Bound">C</a> <a id="4715" href="Categories.Functor.Hom.html#1578" class="Function Operator">][-,-]</a> <a id="4722" class="Symbol">(</a><a id="4723" href="Categories.Category.Core.html#630" class="Function">C.id</a> <a id="4728" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="4730" href="Categories.NaturalTransformation.Core.html#783" class="Function">β.η</a> <a id="4734" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a><a id="4735" class="Symbol">))</a>
|
||
<a id="4756" href="Relation.Binary.Bundles.html#1169" class="Function Operator">≈</a> <a id="4758" href="Function.Construct.Composition.html#5240" class="Function">function</a> <a id="4767" class="Symbol">(</a><a id="4768" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="4771" href="Categories.Functor.Hom.html#1578" class="Function Operator">Hom[</a> <a id="4776" href="Categories.Adjoint.Mate.html#3452" class="Bound">D</a> <a id="4778" href="Categories.Functor.Hom.html#1578" class="Function Operator">][-,-]</a> <a id="4785" class="Symbol">(</a><a id="4786" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="4790" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a> <a id="4792" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="4794" href="Categories.Category.Core.html#630" class="Function">D.id</a><a id="4798" class="Symbol">))</a> <a id="4801" href="Categories.Adjoint.Mate.html#4519" class="Function">D⟶C′</a>
|
||
<a id="4810" href="Categories.Adjoint.Mate.html#4675" class="Function">mate-commute₁</a> <a id="4824" class="Symbol">{</a><a id="4825" href="Categories.Adjoint.Mate.html#4825" class="Bound">f</a><a id="4826" class="Symbol">}</a> <a id="4828" class="Symbol">=</a> <a id="4830" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="4842" href="Categories.NaturalTransformation.Core.html#783" class="Function">β.η</a> <a id="4846" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a> <a id="4848" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="4850" class="Symbol">(</a><a id="4851" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="4854" href="Categories.Adjoint.Mate.html#3458" class="Bound">R′</a> <a id="4857" href="Categories.Adjoint.Mate.html#4825" class="Bound">f</a> <a id="4859" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="4861" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="4874" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a><a id="4875" class="Symbol">)</a> <a id="4877" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="4879" href="Categories.Category.Core.html#630" class="Function">C.id</a> <a id="4885" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4888" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="4896" href="Categories.Category.Core.html#1145" class="Function">identityʳ</a> <a id="4906" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="4914" href="Categories.NaturalTransformation.Core.html#783" class="Function">β.η</a> <a id="4918" href="Categories.Adjoint.Mate.html#3960" class="Bound">Y</a> <a id="4920" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="4922" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="4925" href="Categories.Adjoint.Mate.html#3458" class="Bound">R′</a> <a id="4928" href="Categories.Adjoint.Mate.html#4825" class="Bound">f</a> <a id="4930" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="4932" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="4945" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a> <a id="4957" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4960" href="Categories.Morphism.Reasoning.Core.html#2048" class="Function">pullˡ</a> <a id="4966" class="Symbol">(</a><a id="4967" href="Categories.NaturalTransformation.Core.html#827" class="Function">β.commute</a> <a id="4977" href="Categories.Adjoint.Mate.html#4825" class="Bound">f</a><a id="4978" class="Symbol">)</a> <a id="4980" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="4988" class="Symbol">(</a><a id="4989" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="4992" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a> <a id="4994" href="Categories.Adjoint.Mate.html#4825" class="Bound">f</a> <a id="4996" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="4998" href="Categories.NaturalTransformation.Core.html#783" class="Function">β.η</a> <a id="5002" class="Symbol">(</a><a id="5003" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="5006" href="Categories.Adjoint.Mate.html#3437" class="Bound">L′</a> <a id="5009" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a><a id="5010" class="Symbol">))</a> <a id="5013" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="5015" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="5028" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a> <a id="5031" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="5035" href="Categories.Morphism.Reasoning.Core.html#2347" class="Function">pushʳ</a> <a id="5041" href="Categories.Adjoint.Mate.html#1486" class="Field">commute₁</a> <a id="5050" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="5058" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="5061" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a> <a id="5063" href="Categories.Adjoint.Mate.html#4825" class="Bound">f</a> <a id="5065" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="5067" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="5070" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a> <a id="5072" class="Symbol">(</a><a id="5073" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="5077" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a><a id="5078" class="Symbol">)</a> <a id="5080" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="5082" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="5093" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a> <a id="5101" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="5105" href="Categories.Morphism.Reasoning.Core.html#2485" class="Function">pushˡ</a> <a id="5111" class="Symbol">(</a><a id="5112" href="Categories.Functor.Core.html#565" class="Field">homomorphism</a> <a id="5125" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a><a id="5126" class="Symbol">)</a> <a id="5128" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="5136" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="5139" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a> <a id="5141" class="Symbol">(</a><a id="5142" href="Categories.Adjoint.Mate.html#4825" class="Bound">f</a> <a id="5144" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="5148" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="5152" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a><a id="5153" class="Symbol">)</a> <a id="5155" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="5157" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="5168" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a> <a id="5179" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5182" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="5191" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a> <a id="5193" class="Symbol">(</a><a id="5194" href="Categories.Category.Core.html#3005" class="Function">DH.⟺</a> <a id="5199" href="Categories.Category.Core.html#1096" class="Function">D.identityˡ</a><a id="5210" class="Symbol">)</a> <a id="5212" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="5220" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="5228" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="5231" href="Categories.Adjoint.Mate.html#3456" class="Bound">R</a> <a id="5233" class="Symbol">(</a><a id="5234" href="Categories.Category.Core.html#630" class="Function">D.id</a> <a id="5239" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="5243" href="Categories.Adjoint.Mate.html#4825" class="Bound">f</a> <a id="5245" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="5249" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="5253" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a><a id="5254" class="Symbol">)</a> <a id="5256" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="5258" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="5269" href="Categories.Adjoint.Mate.html#3948" class="Bound">X</a> <a id="5271" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="5276" class="Keyword">module</a> <a id="5283" href="Categories.Adjoint.Mate.html#5283" class="Module">_</a> <a id="5285" class="Symbol">{</a><a id="5286" href="Categories.Adjoint.Mate.html#5286" class="Bound">X</a> <a id="5288" class="Symbol">:</a> <a id="5290" href="Categories.Category.Core.html#559" class="Function">C.Obj</a><a id="5295" class="Symbol">}</a> <a id="5297" class="Symbol">{</a><a id="5298" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a> <a id="5300" class="Symbol">:</a> <a id="5302" href="Categories.Category.Core.html#559" class="Function">D.Obj</a><a id="5307" class="Symbol">}</a> <a id="5309" class="Keyword">where</a>
|
||
<a id="5319" class="Keyword">open</a> <a id="5324" href="Categories.Adjoint.Mate.html#3731" class="Module">D</a> <a id="5326" class="Keyword">hiding</a> <a id="5333" class="Symbol">(</a><a id="5334" href="Categories.Category.Core.html#595" class="Function Operator">_≈_</a><a id="5337" class="Symbol">)</a>
|
||
<a id="5343" class="Keyword">open</a> <a id="5348" href="Categories.Morphism.Reasoning.html" class="Module">MR</a> <a id="5351" href="Categories.Adjoint.Mate.html#3452" class="Bound">D</a>
|
||
<a id="5357" class="Keyword">open</a> <a id="5362" href="Categories.Category.Core.html#2462" class="Module">D.HomReasoning</a>
|
||
<a id="5381" class="Keyword">module</a> <a id="5388" href="Categories.Adjoint.Mate.html#5388" class="Module">CH</a> <a id="5391" class="Symbol">=</a> <a id="5393" href="Categories.Category.Core.html#2462" class="Module">C.HomReasoning</a>
|
||
<a id="5412" class="Keyword">private</a>
|
||
<a id="5426" href="Categories.Adjoint.Mate.html#5426" class="Function">From</a> <a id="5431" class="Symbol">:</a> <a id="5433" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="5440" class="Symbol">_</a> <a id="5442" class="Symbol">_</a>
|
||
<a id="5450" href="Categories.Adjoint.Mate.html#5426" class="Function">From</a> <a id="5455" class="Symbol">=</a> <a id="5457" href="Categories.Functor.Core.html#432" class="Function">L′⊣R′.Hom[-,R-].F₀</a> <a id="5476" class="Symbol">(</a><a id="5477" href="Categories.Adjoint.Mate.html#5286" class="Bound">X</a> <a id="5479" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="5481" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a><a id="5482" class="Symbol">)</a>
|
||
<a id="5490" href="Categories.Adjoint.Mate.html#5490" class="Function">To</a> <a id="5493" class="Symbol">:</a> <a id="5495" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="5502" class="Symbol">_</a> <a id="5504" class="Symbol">_</a>
|
||
<a id="5512" href="Categories.Adjoint.Mate.html#5490" class="Function">To</a> <a id="5515" class="Symbol">=</a> <a id="5517" href="Categories.Functor.Core.html#432" class="Function">L⊣R.Hom[L-,-].F₀</a> <a id="5534" class="Symbol">(</a><a id="5535" href="Categories.Adjoint.Mate.html#5286" class="Bound">X</a> <a id="5537" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="5539" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a><a id="5540" class="Symbol">)</a>
|
||
<a id="5548" class="Keyword">module</a> <a id="5555" href="Categories.Adjoint.Mate.html#5555" class="Module">From</a> <a id="5560" class="Symbol">=</a> <a id="5562" href="Relation.Binary.Bundles.html#1080" class="Module">Setoid</a> <a id="5569" href="Categories.Adjoint.Mate.html#5426" class="Function">From</a>
|
||
<a id="5580" class="Keyword">module</a> <a id="5587" href="Categories.Adjoint.Mate.html#5587" class="Module">To</a> <a id="5590" class="Symbol">=</a> <a id="5592" href="Relation.Binary.Bundles.html#1080" class="Module">Setoid</a> <a id="5599" href="Categories.Adjoint.Mate.html#5490" class="Function">To</a>
|
||
<a id="5608" href="Categories.Adjoint.Mate.html#5608" class="Function">SS</a> <a id="5611" class="Symbol">:</a> <a id="5613" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="5620" class="Symbol">_</a> <a id="5622" class="Symbol">_</a>
|
||
<a id="5630" href="Categories.Adjoint.Mate.html#5608" class="Function">SS</a> <a id="5633" class="Symbol">=</a> <a id="5635" href="Categories.Adjoint.Mate.html#5426" class="Function">From</a> <a id="5640" href="Categories.Adjoint.Mate.html#474" class="Function Operator">⇨</a> <a id="5642" href="Categories.Adjoint.Mate.html#5490" class="Function">To</a>
|
||
<a id="5649" class="Keyword">open</a> <a id="5654" href="Relation.Binary.Bundles.html#1080" class="Module">Setoid</a> <a id="5661" href="Categories.Adjoint.Mate.html#5608" class="Function">SS</a> <a id="5664" class="Keyword">using</a> <a id="5670" class="Symbol">(</a><a id="5671" href="Relation.Binary.Bundles.html#1169" class="Field Operator">_≈_</a><a id="5674" class="Symbol">)</a>
|
||
<a id="5680" class="Keyword">private</a>
|
||
<a id="5694" class="Comment">-- annoyingly the new bundles don't export this</a>
|
||
<a id="5748" href="Categories.Adjoint.Mate.html#5748" class="Function">C⟶D</a> <a id="5752" class="Symbol">:</a> <a id="5754" href="Function.Bundles.html#2043" class="Record">Func</a> <a id="5759" href="Categories.Category.Core.html#1815" class="Function">C.hom-setoid</a> <a id="5772" href="Categories.Category.Core.html#1815" class="Function">D.hom-setoid</a>
|
||
<a id="5791" href="Categories.Adjoint.Mate.html#5748" class="Function">C⟶D</a> <a id="5795" class="Symbol">=</a> <a id="5797" class="Keyword">record</a>
|
||
<a id="5812" class="Symbol">{</a> <a id="5814" href="Function.Bundles.html#2094" class="Field">to</a> <a id="5817" class="Symbol">=</a> <a id="5819" href="Function.Bundles.html#7418" class="Function">L′⊣R′.Hom-inverse.from</a> <a id="5842" class="Symbol">{</a><a id="5843" href="Categories.Adjoint.Mate.html#5286" class="Bound">X</a><a id="5844" class="Symbol">}</a> <a id="5846" class="Symbol">{</a><a id="5847" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a><a id="5848" class="Symbol">}</a>
|
||
<a id="5858" class="Symbol">;</a> <a id="5860" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="5865" class="Symbol">=</a> <a id="5867" href="Function.Bundles.html#7483" class="Function">L′⊣R′.Hom-inverse.from-cong</a>
|
||
<a id="5903" class="Symbol">}</a>
|
||
<a id="5911" href="Categories.Adjoint.Mate.html#5911" class="Function">C⟶D′</a> <a id="5916" class="Symbol">:</a> <a id="5918" href="Function.Bundles.html#2043" class="Record">Func</a> <a id="5923" href="Categories.Category.Core.html#1815" class="Function">C.hom-setoid</a> <a id="5936" href="Categories.Category.Core.html#1815" class="Function">D.hom-setoid</a>
|
||
<a id="5955" href="Categories.Adjoint.Mate.html#5911" class="Function">C⟶D′</a> <a id="5960" class="Symbol">=</a> <a id="5962" class="Keyword">record</a>
|
||
<a id="5977" class="Symbol">{</a> <a id="5979" href="Function.Bundles.html#2094" class="Field">to</a> <a id="5982" class="Symbol">=</a> <a id="5984" href="Function.Bundles.html#7418" class="Function">L⊣R.Hom-inverse.from</a> <a id="6005" class="Symbol">{</a><a id="6006" href="Categories.Adjoint.Mate.html#5286" class="Bound">X</a><a id="6007" class="Symbol">}</a> <a id="6009" class="Symbol">{</a><a id="6010" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a><a id="6011" class="Symbol">}</a>
|
||
<a id="6021" class="Symbol">;</a> <a id="6023" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="6028" class="Symbol">=</a> <a id="6030" href="Function.Bundles.html#7483" class="Function">L⊣R.Hom-inverse.from-cong</a>
|
||
<a id="6064" class="Symbol">}</a>
|
||
|
||
<a id="6071" href="Categories.Adjoint.Mate.html#6071" class="Function">mate-commute₂</a> <a id="6085" class="Symbol">:</a> <a id="6087" href="Function.Construct.Composition.html#5240" class="Function">function</a> <a id="6096" href="Categories.Adjoint.Mate.html#5748" class="Function">C⟶D</a> <a id="6100" class="Symbol">(</a><a id="6101" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6104" href="Categories.Functor.Hom.html#1578" class="Function Operator">Hom[</a> <a id="6109" href="Categories.Adjoint.Mate.html#3452" class="Bound">D</a> <a id="6111" href="Categories.Functor.Hom.html#1578" class="Function Operator">][-,-]</a> <a id="6118" class="Symbol">(</a><a id="6119" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="6123" href="Categories.Adjoint.Mate.html#5286" class="Bound">X</a> <a id="6125" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="6127" href="Categories.Category.Core.html#630" class="Function">D.id</a><a id="6131" class="Symbol">))</a>
|
||
<a id="6152" href="Relation.Binary.Bundles.html#1169" class="Function Operator">≈</a> <a id="6154" href="Function.Construct.Composition.html#5240" class="Function">function</a> <a id="6163" class="Symbol">(</a><a id="6164" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6167" href="Categories.Functor.Hom.html#1578" class="Function Operator">Hom[</a> <a id="6172" href="Categories.Adjoint.Mate.html#3450" class="Bound">C</a> <a id="6174" href="Categories.Functor.Hom.html#1578" class="Function Operator">][-,-]</a> <a id="6181" class="Symbol">(</a><a id="6182" href="Categories.Category.Core.html#630" class="Function">C.id</a> <a id="6187" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="6189" href="Categories.NaturalTransformation.Core.html#783" class="Function">β.η</a> <a id="6193" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a><a id="6194" class="Symbol">))</a> <a id="6197" href="Categories.Adjoint.Mate.html#5911" class="Function">C⟶D′</a>
|
||
<a id="6206" href="Categories.Adjoint.Mate.html#6071" class="Function">mate-commute₂</a> <a id="6220" class="Symbol">{</a><a id="6221" href="Categories.Adjoint.Mate.html#6221" class="Bound">f</a><a id="6222" class="Symbol">}</a> <a id="6224" class="Symbol">=</a> <a id="6226" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="6238" href="Categories.Category.Core.html#630" class="Function">D.id</a> <a id="6243" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6245" class="Symbol">(</a><a id="6246" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="6261" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a> <a id="6263" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6265" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6268" href="Categories.Adjoint.Mate.html#3437" class="Bound">L′</a> <a id="6271" href="Categories.Adjoint.Mate.html#6221" class="Bound">f</a><a id="6272" class="Symbol">)</a> <a id="6274" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6276" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="6280" href="Categories.Adjoint.Mate.html#5286" class="Bound">X</a> <a id="6283" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6286" href="Categories.Category.Core.html#1096" class="Function">identityˡ</a> <a id="6296" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6304" class="Symbol">(</a><a id="6305" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="6320" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a> <a id="6322" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6324" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6327" href="Categories.Adjoint.Mate.html#3437" class="Bound">L′</a> <a id="6330" href="Categories.Adjoint.Mate.html#6221" class="Bound">f</a><a id="6331" class="Symbol">)</a> <a id="6333" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6335" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="6339" href="Categories.Adjoint.Mate.html#5286" class="Bound">X</a> <a id="6349" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="6353" href="Categories.Morphism.Reasoning.Core.html#2347" class="Function">pushʳ</a> <a id="6359" class="Symbol">(</a><a id="6360" href="Categories.NaturalTransformation.Core.html#827" class="Function">α.commute</a> <a id="6370" href="Categories.Adjoint.Mate.html#6221" class="Bound">f</a><a id="6371" class="Symbol">)</a> <a id="6373" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="6381" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="6396" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a> <a id="6398" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6400" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="6404" class="Symbol">(</a><a id="6405" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="6408" href="Categories.Adjoint.Mate.html#3458" class="Bound">R′</a> <a id="6411" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a><a id="6412" class="Symbol">)</a> <a id="6414" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6416" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6419" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a> <a id="6421" href="Categories.Adjoint.Mate.html#6221" class="Bound">f</a> <a id="6426" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="6430" href="Categories.Morphism.Reasoning.Core.html#2485" class="Function">pushˡ</a> <a id="6436" href="Categories.Adjoint.Mate.html#1548" class="Field">commute₂</a> <a id="6445" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="6453" class="Symbol">(</a><a id="6454" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="6467" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a> <a id="6469" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6471" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6474" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a> <a id="6476" class="Symbol">(</a><a id="6477" href="Categories.NaturalTransformation.Core.html#783" class="Function">β.η</a> <a id="6481" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a><a id="6482" class="Symbol">))</a> <a id="6485" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6487" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6490" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a> <a id="6492" href="Categories.Adjoint.Mate.html#6221" class="Bound">f</a> <a id="6498" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="6502" href="Categories.Morphism.Reasoning.Core.html#2347" class="Function">pushʳ</a> <a id="6508" class="Symbol">(</a><a id="6509" href="Categories.Functor.Core.html#565" class="Field">homomorphism</a> <a id="6522" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a><a id="6523" class="Symbol">)</a> <a id="6525" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="6533" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="6546" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a> <a id="6548" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6550" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6553" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a> <a id="6555" class="Symbol">(</a><a id="6556" href="Categories.NaturalTransformation.Core.html#783" class="Function">β.η</a> <a id="6560" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a> <a id="6562" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="6566" href="Categories.Adjoint.Mate.html#6221" class="Bound">f</a><a id="6567" class="Symbol">)</a> <a id="6578" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6581" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="6589" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="6598" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a> <a id="6600" class="Symbol">(</a><a id="6601" href="Categories.Category.Core.html#1706" class="Function">C.∘-resp-≈ʳ</a> <a id="6613" class="Symbol">(</a><a id="6614" href="Categories.Category.Core.html#3005" class="Function">CH.⟺</a> <a id="6619" href="Categories.Category.Core.html#1145" class="Function">C.identityʳ</a><a id="6630" class="Symbol">))</a> <a id="6633" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6641" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="6654" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a> <a id="6656" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="6658" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6661" href="Categories.Adjoint.Mate.html#3435" class="Bound">L</a> <a id="6663" class="Symbol">(</a><a id="6664" href="Categories.NaturalTransformation.Core.html#783" class="Function">β.η</a> <a id="6668" href="Categories.Adjoint.Mate.html#5298" class="Bound">Y</a> <a id="6670" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="6674" href="Categories.Adjoint.Mate.html#6221" class="Bound">f</a> <a id="6676" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="6680" href="Categories.Category.Core.html#630" class="Function">C.id</a><a id="6684" class="Symbol">)</a> <a id="6686" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="6689" class="Comment">-- alternatively, if commute₃ and commute₄ are shown, then a Mate can be constructed.</a>
|
||
|
||
<a id="6776" class="Keyword">module</a> <a id="6783" href="Categories.Adjoint.Mate.html#6783" class="Module">_</a> <a id="6785" class="Symbol">{</a><a id="6786" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="6788" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="6791" class="Symbol">:</a> <a id="6793" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="6801" href="Categories.Adjoint.Mate.html#936" class="Generalizable">C</a> <a id="6803" href="Categories.Adjoint.Mate.html#938" class="Generalizable">D</a><a id="6804" class="Symbol">}</a> <a id="6806" class="Symbol">{</a><a id="6807" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="6809" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="6812" class="Symbol">:</a> <a id="6814" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="6822" href="Categories.Adjoint.Mate.html#938" class="Generalizable">D</a> <a id="6824" href="Categories.Adjoint.Mate.html#936" class="Generalizable">C</a><a id="6825" class="Symbol">}</a>
|
||
<a id="6836" class="Symbol">{</a><a id="6837" href="Categories.Adjoint.Mate.html#6837" class="Bound">L⊣R</a> <a id="6841" class="Symbol">:</a> <a id="6843" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="6845" href="Categories.Adjoint.html#7818" class="Function Operator">⊣</a> <a id="6847" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a><a id="6848" class="Symbol">}</a> <a id="6850" class="Symbol">{</a><a id="6851" href="Categories.Adjoint.Mate.html#6851" class="Bound">L′⊣R′</a> <a id="6857" class="Symbol">:</a> <a id="6859" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="6862" href="Categories.Adjoint.html#7818" class="Function Operator">⊣</a> <a id="6864" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a><a id="6866" class="Symbol">}</a>
|
||
<a id="6877" class="Symbol">{</a><a id="6878" href="Categories.Adjoint.Mate.html#6878" class="Bound">α</a> <a id="6880" class="Symbol">:</a> <a id="6882" href="Categories.NaturalTransformation.Core.html#466" class="Record">NaturalTransformation</a> <a id="6904" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="6906" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a><a id="6908" class="Symbol">}</a>
|
||
<a id="6919" class="Symbol">{</a><a id="6920" href="Categories.Adjoint.Mate.html#6920" class="Bound">β</a> <a id="6922" class="Symbol">:</a> <a id="6924" href="Categories.NaturalTransformation.Core.html#466" class="Record">NaturalTransformation</a> <a id="6946" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="6949" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a><a id="6950" class="Symbol">}</a> <a id="6952" class="Keyword">where</a>
|
||
<a id="6960" class="Keyword">private</a>
|
||
<a id="6972" class="Keyword">open</a> <a id="6977" href="Categories.Functor.Core.html#248" class="Module">Functor</a>
|
||
<a id="6989" class="Keyword">module</a> <a id="6996" href="Categories.Adjoint.Mate.html#6996" class="Module">C</a> <a id="7002" class="Symbol">=</a> <a id="7004" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="7013" href="Categories.Adjoint.Mate.html#6801" class="Bound">C</a>
|
||
<a id="7019" class="Keyword">module</a> <a id="7026" href="Categories.Adjoint.Mate.html#7026" class="Module">D</a> <a id="7032" class="Symbol">=</a> <a id="7034" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="7043" href="Categories.Adjoint.Mate.html#6803" class="Bound">D</a>
|
||
<a id="7049" class="Keyword">module</a> <a id="7056" href="Categories.Adjoint.Mate.html#7056" class="Module">α</a> <a id="7062" class="Symbol">=</a> <a id="7064" href="Categories.NaturalTransformation.Core.html#466" class="Module">NaturalTransformation</a> <a id="7086" href="Categories.Adjoint.Mate.html#6878" class="Bound">α</a>
|
||
<a id="7092" class="Keyword">module</a> <a id="7099" href="Categories.Adjoint.Mate.html#7099" class="Module">β</a> <a id="7105" class="Symbol">=</a> <a id="7107" href="Categories.NaturalTransformation.Core.html#466" class="Module">NaturalTransformation</a> <a id="7129" href="Categories.Adjoint.Mate.html#6920" class="Bound">β</a>
|
||
<a id="7135" class="Keyword">module</a> <a id="7142" href="Categories.Adjoint.Mate.html#7142" class="Module">L⊣R</a> <a id="7148" class="Symbol">=</a> <a id="7150" href="Categories.Adjoint.html#1260" class="Module">Adjoint</a> <a id="7158" href="Categories.Adjoint.Mate.html#6837" class="Bound">L⊣R</a>
|
||
<a id="7166" class="Keyword">module</a> <a id="7173" href="Categories.Adjoint.Mate.html#7173" class="Module">L′⊣R′</a> <a id="7179" class="Symbol">=</a> <a id="7181" href="Categories.Adjoint.html#1260" class="Module">Adjoint</a> <a id="7189" href="Categories.Adjoint.Mate.html#6851" class="Bound">L′⊣R′</a>
|
||
|
||
<a id="7198" class="Keyword">module</a> <a id="7205" href="Categories.Adjoint.Mate.html#7205" class="Module">_</a> <a id="7207" class="Symbol">(</a><a id="7208" href="Categories.Adjoint.Mate.html#7208" class="Bound">commute₃</a> <a id="7217" class="Symbol">:</a> <a id="7219" class="Symbol">∀</a> <a id="7221" class="Symbol">{</a><a id="7222" href="Categories.Adjoint.Mate.html#7222" class="Bound">X</a><a id="7223" class="Symbol">}</a> <a id="7225" class="Symbol">→</a> <a id="7227" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="7240" class="Symbol">(</a><a id="7241" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7244" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="7247" href="Categories.Adjoint.Mate.html#7222" class="Bound">X</a><a id="7248" class="Symbol">)</a> <a id="7250" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="7254" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7257" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="7259" class="Symbol">(</a><a id="7260" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="7264" class="Symbol">(</a><a id="7265" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7268" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="7271" href="Categories.Adjoint.Mate.html#7222" class="Bound">X</a><a id="7272" class="Symbol">))</a> <a id="7275" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="7279" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7282" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="7284" class="Symbol">(</a><a id="7285" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="7298" href="Categories.Adjoint.Mate.html#7222" class="Bound">X</a><a id="7299" class="Symbol">)</a> <a id="7301" href="Categories.Category.Core.html#595" class="Function Operator">D.≈</a> <a id="7305" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="7309" href="Categories.Adjoint.Mate.html#7222" class="Bound">X</a><a id="7310" class="Symbol">)</a> <a id="7312" class="Keyword">where</a>
|
||
<a id="7322" class="Keyword">open</a> <a id="7327" href="Categories.Adjoint.Mate.html#6996" class="Module">C</a>
|
||
<a id="7333" class="Keyword">open</a> <a id="7338" href="Categories.Category.Core.html#2462" class="Module">HomReasoning</a>
|
||
|
||
<a id="7356" href="Categories.Adjoint.Mate.html#7356" class="Function">commute₁</a> <a id="7365" class="Symbol">:</a> <a id="7367" class="Symbol">∀</a> <a id="7369" class="Symbol">{</a><a id="7370" href="Categories.Adjoint.Mate.html#7370" class="Bound">X</a><a id="7371" class="Symbol">}</a> <a id="7373" class="Symbol">→</a> <a id="7375" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7378" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="7380" class="Symbol">(</a><a id="7381" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="7385" href="Categories.Adjoint.Mate.html#7370" class="Bound">X</a><a id="7386" class="Symbol">)</a> <a id="7388" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="7390" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="7401" href="Categories.Adjoint.Mate.html#7370" class="Bound">X</a> <a id="7403" href="Categories.Category.Core.html#595" class="Function Operator">≈</a> <a id="7405" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="7409" class="Symbol">(</a><a id="7410" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7413" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="7416" href="Categories.Adjoint.Mate.html#7370" class="Bound">X</a><a id="7417" class="Symbol">)</a> <a id="7419" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="7421" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="7434" href="Categories.Adjoint.Mate.html#7370" class="Bound">X</a>
|
||
<a id="7440" href="Categories.Adjoint.Mate.html#7356" class="Function">commute₁</a> <a id="7449" class="Symbol">{</a><a id="7450" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a><a id="7451" class="Symbol">}</a> <a id="7453" class="Symbol">=</a> <a id="7455" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="7467" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7470" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="7472" class="Symbol">(</a><a id="7473" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="7477" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a><a id="7478" class="Symbol">)</a> <a id="7480" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="7482" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="7493" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a>
|
||
<a id="7503" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="7507" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="7516" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="7518" href="Categories.Adjoint.Mate.html#7208" class="Bound">commute₃</a> <a id="7527" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="7535" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="7543" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7546" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="7548" class="Symbol">(</a><a id="7549" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="7562" class="Symbol">(</a><a id="7563" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7566" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="7569" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a><a id="7570" class="Symbol">)</a> <a id="7572" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="7576" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7579" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="7581" class="Symbol">(</a><a id="7582" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="7586" class="Symbol">(</a><a id="7587" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7590" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="7593" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a><a id="7594" class="Symbol">))</a> <a id="7597" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="7601" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7604" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="7606" class="Symbol">(</a><a id="7607" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="7620" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a><a id="7621" class="Symbol">))</a> <a id="7624" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="7626" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="7637" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a>
|
||
<a id="7647" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="7651" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="7660" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="7662" class="Symbol">(</a><a id="7663" href="Categories.Category.Core.html#1706" class="Function">D.∘-resp-≈ʳ</a> <a id="7675" class="Symbol">(</a><a id="7676" href="Categories.Functor.Core.html#565" class="Field">homomorphism</a> <a id="7689" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a><a id="7690" class="Symbol">))</a> <a id="7693" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="7701" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="7709" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7712" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="7714" class="Symbol">(</a><a id="7715" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="7728" class="Symbol">(</a><a id="7729" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7732" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="7735" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a><a id="7736" class="Symbol">)</a> <a id="7738" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="7742" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7745" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="7747" class="Symbol">(</a><a id="7748" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="7752" class="Symbol">(</a><a id="7753" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7756" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="7759" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a><a id="7760" class="Symbol">)</a> <a id="7762" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="7764" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="7777" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a><a id="7778" class="Symbol">))</a> <a id="7781" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="7783" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="7794" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a>
|
||
<a id="7804" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="7807" href="Categories.Adjoint.html#2533" class="Function">L⊣R.LRadjunct≈id</a> <a id="7824" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="7832" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="7836" class="Symbol">(</a><a id="7837" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7840" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="7843" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a><a id="7844" class="Symbol">)</a> <a id="7846" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="7848" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="7861" href="Categories.Adjoint.Mate.html#7450" class="Bound">X</a>
|
||
<a id="7871" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="7876" class="Keyword">module</a> <a id="7883" href="Categories.Adjoint.Mate.html#7883" class="Module">_</a> <a id="7885" class="Symbol">(</a><a id="7886" href="Categories.Adjoint.Mate.html#7886" class="Bound">commute₄</a> <a id="7895" class="Symbol">:</a> <a id="7897" class="Symbol">∀</a> <a id="7899" class="Symbol">{</a><a id="7900" href="Categories.Adjoint.Mate.html#7900" class="Bound">X</a><a id="7901" class="Symbol">}</a> <a id="7903" class="Symbol">→</a> <a id="7905" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7908" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="7910" class="Symbol">(</a><a id="7911" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="7926" href="Categories.Adjoint.Mate.html#7900" class="Bound">X</a><a id="7927" class="Symbol">)</a> <a id="7929" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="7933" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="7936" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="7938" class="Symbol">(</a><a id="7939" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="7943" class="Symbol">(</a><a id="7944" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7947" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="7950" href="Categories.Adjoint.Mate.html#7900" class="Bound">X</a><a id="7951" class="Symbol">))</a> <a id="7954" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="7958" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="7969" class="Symbol">(</a><a id="7970" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="7973" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="7976" href="Categories.Adjoint.Mate.html#7900" class="Bound">X</a><a id="7977" class="Symbol">)</a> <a id="7979" href="Categories.Category.Core.html#595" class="Function Operator">C.≈</a> <a id="7983" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="7987" href="Categories.Adjoint.Mate.html#7900" class="Bound">X</a><a id="7988" class="Symbol">)</a> <a id="7990" class="Keyword">where</a>
|
||
<a id="8000" class="Keyword">open</a> <a id="8005" href="Categories.Adjoint.Mate.html#7026" class="Module">D</a>
|
||
<a id="8011" class="Keyword">open</a> <a id="8016" href="Categories.Category.Core.html#2462" class="Module">HomReasoning</a>
|
||
<a id="8033" class="Keyword">open</a> <a id="8038" href="Categories.Morphism.Reasoning.html" class="Module">MR</a> <a id="8041" href="Categories.Adjoint.Mate.html#6801" class="Bound">C</a>
|
||
|
||
<a id="8048" href="Categories.Adjoint.Mate.html#8048" class="Function">commute₂</a> <a id="8057" class="Symbol">:</a> <a id="8059" class="Symbol">∀</a> <a id="8061" class="Symbol">{</a><a id="8062" href="Categories.Adjoint.Mate.html#8062" class="Bound">X</a><a id="8063" class="Symbol">}</a> <a id="8065" class="Symbol">→</a> <a id="8067" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="8080" href="Categories.Adjoint.Mate.html#8062" class="Bound">X</a> <a id="8082" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="8084" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8087" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="8089" class="Symbol">(</a><a id="8090" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="8094" href="Categories.Adjoint.Mate.html#8062" class="Bound">X</a><a id="8095" class="Symbol">)</a> <a id="8097" href="Categories.Category.Core.html#595" class="Function Operator">≈</a> <a id="8099" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="8114" href="Categories.Adjoint.Mate.html#8062" class="Bound">X</a> <a id="8116" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="8118" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="8122" class="Symbol">(</a><a id="8123" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8126" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="8129" href="Categories.Adjoint.Mate.html#8062" class="Bound">X</a><a id="8130" class="Symbol">)</a>
|
||
<a id="8136" href="Categories.Adjoint.Mate.html#8048" class="Function">commute₂</a> <a id="8145" class="Symbol">{</a><a id="8146" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a><a id="8147" class="Symbol">}</a> <a id="8149" class="Symbol">=</a> <a id="8151" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="8163" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="8176" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a> <a id="8178" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="8180" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8183" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="8185" class="Symbol">(</a><a id="8186" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="8190" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a><a id="8191" class="Symbol">)</a>
|
||
<a id="8201" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="8205" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="8213" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="8222" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="8224" href="Categories.Adjoint.Mate.html#7886" class="Bound">commute₄</a> <a id="8233" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="8241" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="8254" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a> <a id="8256" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="8258" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8261" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="8263" class="Symbol">(</a><a id="8264" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8267" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="8269" class="Symbol">(</a><a id="8270" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="8285" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a><a id="8286" class="Symbol">)</a> <a id="8288" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="8292" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8295" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="8297" class="Symbol">(</a><a id="8298" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="8302" class="Symbol">(</a><a id="8303" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8306" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="8309" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a><a id="8310" class="Symbol">))</a> <a id="8313" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="8317" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="8328" class="Symbol">(</a><a id="8329" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8332" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="8335" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a><a id="8336" class="Symbol">))</a>
|
||
<a id="8347" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="8351" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="8359" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="8368" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="8370" class="Symbol">(</a><a id="8371" href="Categories.Morphism.Reasoning.Core.html#2485" class="Function">pushˡ</a> <a id="8377" class="Symbol">(</a><a id="8378" href="Categories.Functor.Core.html#565" class="Field">homomorphism</a> <a id="8391" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a><a id="8392" class="Symbol">))</a> <a id="8395" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="8403" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="8416" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a> <a id="8418" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="8420" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8423" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="8425" class="Symbol">(</a><a id="8426" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8429" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="8431" class="Symbol">(</a><a id="8432" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="8447" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a> <a id="8449" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="8451" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="8455" class="Symbol">(</a><a id="8456" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8459" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="8462" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a><a id="8463" class="Symbol">))</a> <a id="8466" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="8470" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="8481" class="Symbol">(</a><a id="8482" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8485" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="8488" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a><a id="8489" class="Symbol">))</a>
|
||
<a id="8500" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8503" href="Categories.Adjoint.html#2026" class="Function">L⊣R.RLadjunct≈id</a> <a id="8520" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8528" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="8543" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a> <a id="8545" href="Categories.Category.Core.html#656" class="Function Operator">∘</a> <a id="8547" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="8551" class="Symbol">(</a><a id="8552" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8555" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="8558" href="Categories.Adjoint.Mate.html#8146" class="Bound">X</a><a id="8559" class="Symbol">)</a>
|
||
<a id="8569" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
|
||
<a id="8575" href="Categories.Adjoint.Mate.html#8575" class="Function">mate′</a> <a id="8581" class="Symbol">:</a> <a id="8583" class="Symbol">(∀</a> <a id="8586" class="Symbol">{</a><a id="8587" href="Categories.Adjoint.Mate.html#8587" class="Bound">X</a><a id="8588" class="Symbol">}</a> <a id="8590" class="Symbol">→</a> <a id="8592" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.counit.η</a> <a id="8605" class="Symbol">(</a><a id="8606" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8609" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="8612" href="Categories.Adjoint.Mate.html#8587" class="Bound">X</a><a id="8613" class="Symbol">)</a> <a id="8615" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="8619" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8622" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="8624" class="Symbol">(</a><a id="8625" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="8629" class="Symbol">(</a><a id="8630" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8633" href="Categories.Adjoint.Mate.html#6788" class="Bound">L′</a> <a id="8636" href="Categories.Adjoint.Mate.html#8587" class="Bound">X</a><a id="8637" class="Symbol">))</a> <a id="8640" href="Categories.Category.Core.html#656" class="Function Operator">D.∘</a> <a id="8644" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8647" href="Categories.Adjoint.Mate.html#6786" class="Bound">L</a> <a id="8649" class="Symbol">(</a><a id="8650" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.unit.η</a> <a id="8663" href="Categories.Adjoint.Mate.html#8587" class="Bound">X</a><a id="8664" class="Symbol">)</a> <a id="8666" href="Categories.Category.Core.html#595" class="Function Operator">D.≈</a> <a id="8670" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="8674" href="Categories.Adjoint.Mate.html#8587" class="Bound">X</a><a id="8675" class="Symbol">)</a> <a id="8677" class="Symbol">→</a>
|
||
<a id="8689" class="Symbol">(∀</a> <a id="8692" class="Symbol">{</a><a id="8693" href="Categories.Adjoint.Mate.html#8693" class="Bound">X</a><a id="8694" class="Symbol">}</a> <a id="8696" class="Symbol">→</a> <a id="8698" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8701" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="8703" class="Symbol">(</a><a id="8704" href="Categories.NaturalTransformation.Core.html#783" class="Function">L′⊣R′.counit.η</a> <a id="8719" href="Categories.Adjoint.Mate.html#8693" class="Bound">X</a><a id="8720" class="Symbol">)</a> <a id="8722" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="8726" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="8729" href="Categories.Adjoint.Mate.html#6807" class="Bound">R</a> <a id="8731" class="Symbol">(</a><a id="8732" href="Categories.NaturalTransformation.Core.html#783" class="Function">α.η</a> <a id="8736" class="Symbol">(</a><a id="8737" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8740" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="8743" href="Categories.Adjoint.Mate.html#8693" class="Bound">X</a><a id="8744" class="Symbol">))</a> <a id="8747" href="Categories.Category.Core.html#656" class="Function Operator">C.∘</a> <a id="8751" href="Categories.NaturalTransformation.Core.html#783" class="Function">L⊣R.unit.η</a> <a id="8762" class="Symbol">(</a><a id="8763" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="8766" href="Categories.Adjoint.Mate.html#6809" class="Bound">R′</a> <a id="8769" href="Categories.Adjoint.Mate.html#8693" class="Bound">X</a><a id="8770" class="Symbol">)</a> <a id="8772" href="Categories.Category.Core.html#595" class="Function Operator">C.≈</a> <a id="8776" href="Categories.NaturalTransformation.Core.html#783" class="Field">β.η</a> <a id="8780" href="Categories.Adjoint.Mate.html#8693" class="Bound">X</a><a id="8781" class="Symbol">)</a> <a id="8783" class="Symbol">→</a>
|
||
<a id="8795" href="Categories.Adjoint.Mate.html#1140" class="Record">Mate</a> <a id="8800" href="Categories.Adjoint.Mate.html#6837" class="Bound">L⊣R</a> <a id="8804" href="Categories.Adjoint.Mate.html#6851" class="Bound">L′⊣R′</a> <a id="8810" href="Categories.Adjoint.Mate.html#6878" class="Bound">α</a> <a id="8812" href="Categories.Adjoint.Mate.html#6920" class="Bound">β</a>
|
||
<a id="8816" href="Categories.Adjoint.Mate.html#8575" class="Function">mate′</a> <a id="8822" href="Categories.Adjoint.Mate.html#8822" class="Bound">commute₃</a> <a id="8831" href="Categories.Adjoint.Mate.html#8831" class="Bound">commute₄</a> <a id="8840" class="Symbol">=</a> <a id="8842" class="Keyword">record</a>
|
||
<a id="8853" class="Symbol">{</a> <a id="8855" href="Categories.Adjoint.Mate.html#1486" class="Field">commute₁</a> <a id="8864" class="Symbol">=</a> <a id="8866" href="Categories.Adjoint.Mate.html#7356" class="Function">commute₁</a> <a id="8875" href="Categories.Adjoint.Mate.html#8822" class="Bound">commute₃</a>
|
||
<a id="8888" class="Symbol">;</a> <a id="8890" href="Categories.Adjoint.Mate.html#1548" class="Field">commute₂</a> <a id="8899" class="Symbol">=</a> <a id="8901" href="Categories.Adjoint.Mate.html#8048" class="Function">commute₂</a> <a id="8910" href="Categories.Adjoint.Mate.html#8831" class="Bound">commute₄</a>
|
||
<a id="8923" class="Symbol">}</a>
|
||
</pre></body></html> |