bsc-leon-vatthauer/agda/bsc-thesis/Categories.Category.CartesianClosed.html
2024-02-09 17:53:52 +01:00

293 lines
No EOL
164 KiB
HTML
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

<!DOCTYPE HTML>
<html><head><meta charset="utf-8"><title>Categories.Category.CartesianClosed</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="36" class="Keyword">open</a> <a id="41" class="Keyword">import</a> <a id="48" href="Categories.Category.html" class="Module">Categories.Category</a>
<a id="69" class="Keyword">module</a> <a id="76" href="Categories.Category.CartesianClosed.html" class="Module">Categories.Category.CartesianClosed</a> <a id="112" class="Symbol">{</a><a id="113" href="Categories.Category.CartesianClosed.html#113" class="Bound">o</a> <a id="115" href="Categories.Category.CartesianClosed.html#115" class="Bound"></a> <a id="117" href="Categories.Category.CartesianClosed.html#117" class="Bound">e</a><a id="118" class="Symbol">}</a> <a id="120" class="Symbol">(</a><a id="121" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a> <a id="123" class="Symbol">:</a> <a id="125" href="Categories.Category.Core.html#442" class="Record">Category</a> <a id="134" href="Categories.Category.CartesianClosed.html#113" class="Bound">o</a> <a id="136" href="Categories.Category.CartesianClosed.html#115" class="Bound"></a> <a id="138" href="Categories.Category.CartesianClosed.html#117" class="Bound">e</a><a id="139" class="Symbol">)</a> <a id="141" class="Keyword">where</a>
<a id="148" class="Keyword">open</a> <a id="153" class="Keyword">import</a> <a id="160" href="Level.html" class="Module">Level</a>
<a id="166" class="Keyword">open</a> <a id="171" class="Keyword">import</a> <a id="178" href="Function.html" class="Module">Function</a> <a id="187" class="Keyword">using</a> <a id="193" class="Symbol">(</a><a id="194" href="Function.Base.html#1974" class="Function Operator">_$_</a><a id="197" class="Symbol">;</a> <a id="199" href="Function.Base.html#1638" class="Function">flip</a><a id="203" class="Symbol">)</a>
<a id="205" class="Keyword">open</a> <a id="210" class="Keyword">import</a> <a id="217" href="Data.Product.html" class="Module">Data.Product</a> <a id="230" class="Keyword">using</a> <a id="236" class="Symbol">(</a><a id="237" href="Agda.Builtin.Sigma.html#165" class="Record">Σ</a><a id="238" class="Symbol">;</a> <a id="240" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">_,_</a><a id="243" class="Symbol">;</a> <a id="245" href="Data.Product.Base.html#3109" class="Function">uncurry</a><a id="252" class="Symbol">)</a>
<a id="255" class="Keyword">open</a> <a id="260" class="Keyword">import</a> <a id="267" href="Categories.Category.BinaryProducts.html" class="Module">Categories.Category.BinaryProducts</a> <a id="302" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="304" class="Keyword">open</a> <a id="309" class="Keyword">import</a> <a id="316" href="Categories.Category.Cartesian.html" class="Module">Categories.Category.Cartesian</a> <a id="346" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="348" class="Keyword">open</a> <a id="353" class="Keyword">import</a> <a id="360" href="Categories.Category.Cartesian.Monoidal.html" class="Module">Categories.Category.Cartesian.Monoidal</a> <a id="399" class="Keyword">using</a> <a id="405" class="Symbol">(</a><a id="406" class="Keyword">module</a> <a id="413" href="Categories.Category.Cartesian.Monoidal.html#1139" class="Module">CartesianMonoidal</a><a id="430" class="Symbol">)</a>
<a id="432" class="Keyword">open</a> <a id="437" class="Keyword">import</a> <a id="444" href="Categories.Category.Monoidal.Closed.html" class="Module">Categories.Category.Monoidal.Closed</a> <a id="480" class="Keyword">using</a> <a id="486" class="Symbol">(</a><a id="487" href="Categories.Category.Monoidal.Closed.html#1618" class="Record">Closed</a><a id="493" class="Symbol">)</a>
<a id="495" class="Keyword">open</a> <a id="500" class="Keyword">import</a> <a id="507" href="Categories.Functor.html" class="Module">Categories.Functor</a> <a id="526" class="Keyword">renaming</a> <a id="535" class="Symbol">(</a><a id="536" href="Categories.Functor.html#349" class="Function">id</a> <a id="539" class="Symbol">to</a> <a id="542" class="Function">idF</a><a id="545" class="Symbol">)</a>
<a id="547" class="Keyword">open</a> <a id="552" class="Keyword">import</a> <a id="559" href="Categories.Functor.Bifunctor.html" class="Module">Categories.Functor.Bifunctor</a>
<a id="588" class="Keyword">open</a> <a id="593" class="Keyword">import</a> <a id="600" href="Categories.NaturalTransformation.html" class="Module">Categories.NaturalTransformation</a> <a id="633" class="Keyword">hiding</a> <a id="640" class="Symbol">(</a><a id="641" href="Categories.NaturalTransformation.Core.html#2132" class="Function">id</a><a id="643" class="Symbol">)</a>
<a id="645" class="Keyword">open</a> <a id="650" class="Keyword">import</a> <a id="657" href="Categories.NaturalTransformation.Properties.html" class="Module">Categories.NaturalTransformation.Properties</a>
<a id="701" class="Keyword">open</a> <a id="706" class="Keyword">import</a> <a id="713" href="Categories.Object.Product.html" class="Module">Categories.Object.Product</a> <a id="739" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="743" class="Keyword">hiding</a> <a id="750" class="Symbol">(</a><a id="751" href="Categories.Object.Product.Core.html#2384" class="Function">repack≡id</a><a id="760" class="Symbol">;</a> <a id="762" href="Categories.Object.Product.Core.html#2168" class="Function">repack∘</a><a id="769" class="Symbol">;</a> <a id="771" href="Categories.Object.Product.Core.html#2451" class="Function">repack-cancel</a><a id="784" class="Symbol">;</a> <a id="786" href="Categories.Object.Product.Core.html#2581" class="Function">up-to-iso</a><a id="795" class="Symbol">;</a> <a id="797" href="Categories.Object.Product.Core.html#2819" class="Function">transport-by-iso</a><a id="813" class="Symbol">)</a>
<a id="815" class="Keyword">open</a> <a id="820" class="Keyword">import</a> <a id="827" href="Categories.Object.Exponential.html" class="Module">Categories.Object.Exponential</a> <a id="857" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a> <a id="859" class="Keyword">hiding</a> <a id="866" class="Symbol">(</a><a id="867" href="Categories.Object.Exponential.html#4450" class="Function">repack</a><a id="873" class="Symbol">)</a>
<a id="875" class="Keyword">open</a> <a id="880" class="Keyword">import</a> <a id="887" href="Categories.Object.Terminal.html" class="Module">Categories.Object.Terminal</a> <a id="914" class="Keyword">using</a> <a id="920" class="Symbol">(</a><a id="921" href="Categories.Object.Terminal.html#860" class="Record">Terminal</a><a id="929" class="Symbol">)</a>
<a id="931" class="Keyword">open</a> <a id="936" class="Keyword">import</a> <a id="943" href="Categories.Morphism.html" class="Module">Categories.Morphism</a> <a id="963" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="965" class="Keyword">open</a> <a id="970" class="Keyword">import</a> <a id="977" href="Categories.Morphism.Reasoning.html" class="Module">Categories.Morphism.Reasoning</a> <a id="1007" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="1010" class="Keyword">private</a>
<a id="1020" class="Keyword">module</a> <a id="𝒞"></a><a id="1027" href="Categories.Category.CartesianClosed.html#1027" class="Module">𝒞</a> <a id="1029" class="Symbol">=</a> <a id="1031" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="1040" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="1044" class="Keyword">open</a> <a id="1049" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="1058" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="1062" class="Keyword">open</a> <a id="1067" href="Categories.Category.Core.html#2462" class="Module">HomReasoning</a>
<a id="1082" class="Keyword">open</a> <a id="1087" href="Categories.Category.Core.html#1530" class="Module">Equiv</a>
<a id="1095" class="Keyword">variable</a>
<a id="1108" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="1110" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a> <a id="1112" href="Categories.Category.CartesianClosed.html#1112" class="Generalizable">C</a> <a id="1116" class="Symbol">:</a> <a id="1118" href="Categories.Category.Core.html#559" class="Field">Obj</a>
<a id="1126" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="1128" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="1130" href="Categories.Category.CartesianClosed.html#1130" class="Generalizable">h</a> <a id="1132" href="Categories.Category.CartesianClosed.html#1132" class="Generalizable">i</a> <a id="1134" class="Symbol">:</a> <a id="1136" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="1138" href="Categories.Category.Core.html#575" class="Field Operator"></a> <a id="1140" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a>
<a id="1143" class="Comment">-- Cartesian closed category</a>
<a id="1172" class="Comment">-- is a category with all products and exponentials</a>
<a id="1226" class="Keyword">record</a> <a id="CartesianClosed"></a><a id="1233" href="Categories.Category.CartesianClosed.html#1233" class="Record">CartesianClosed</a> <a id="1249" class="Symbol">:</a> <a id="1251" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1255" class="Symbol">(</a><a id="1256" href="Level.html#602" class="Function">levelOfTerm</a> <a id="1268" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a><a id="1269" class="Symbol">)</a> <a id="1271" class="Keyword">where</a>
<a id="1279" class="Keyword">infixr</a> <a id="1286" class="Number">9</a> <a id="1288" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">_^_</a>
<a id="1294" class="Comment">-- an alternative notation for exponential, which emphasizes its internal hom natural</a>
<a id="1382" class="Keyword">infixr</a> <a id="1389" class="Number">5</a> <a id="1391" href="Categories.Category.CartesianClosed.html#1564" class="Function Operator">_⇨_</a>
<a id="1398" class="Keyword">field</a>
<a id="CartesianClosed.cartesian"></a><a id="1408" href="Categories.Category.CartesianClosed.html#1408" class="Field">cartesian</a> <a id="1418" class="Symbol">:</a> <a id="1420" href="Categories.Category.Cartesian.html#727" class="Record">Cartesian</a>
<a id="CartesianClosed.exp"></a><a id="1434" href="Categories.Category.CartesianClosed.html#1434" class="Field">exp</a> <a id="1444" class="Symbol">:</a> <a id="1446" href="Categories.Object.Exponential.html#583" class="Record">Exponential</a> <a id="1458" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="1460" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a>
<a id="1465" class="Keyword">module</a> <a id="CartesianClosed.exp"></a><a id="1472" href="Categories.Category.CartesianClosed.html#1472" class="Module">exp</a> <a id="1476" class="Symbol">{</a><a id="1477" href="Categories.Category.CartesianClosed.html#1477" class="Bound">A</a> <a id="1479" href="Categories.Category.CartesianClosed.html#1479" class="Bound">B</a><a id="1480" class="Symbol">}</a> <a id="1482" class="Symbol">=</a> <a id="1484" href="Categories.Object.Exponential.html#583" class="Module">Exponential</a> <a id="1496" class="Symbol">(</a><a id="1497" href="Categories.Category.CartesianClosed.html#1434" class="Field">exp</a> <a id="1501" class="Symbol">{</a><a id="1502" href="Categories.Category.CartesianClosed.html#1477" class="Bound">A</a><a id="1503" class="Symbol">}</a> <a id="1505" class="Symbol">{</a><a id="1506" href="Categories.Category.CartesianClosed.html#1479" class="Bound">B</a><a id="1507" class="Symbol">})</a>
<a id="CartesianClosed._^_"></a><a id="1513" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">_^_</a> <a id="1517" class="Symbol">:</a> <a id="1519" href="Categories.Category.Core.html#559" class="Field">Obj</a> <a id="1523" class="Symbol"></a> <a id="1525" href="Categories.Category.Core.html#559" class="Field">Obj</a> <a id="1529" class="Symbol"></a> <a id="1531" href="Categories.Category.Core.html#559" class="Field">Obj</a>
<a id="1537" href="Categories.Category.CartesianClosed.html#1537" class="Bound">B</a> <a id="1539" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">^</a> <a id="1541" href="Categories.Category.CartesianClosed.html#1541" class="Bound">A</a> <a id="1543" class="Symbol">=</a> <a id="1545" href="Categories.Object.Exponential.html#643" class="Function">exp.B^A</a> <a id="1553" class="Symbol">{</a><a id="1554" href="Categories.Category.CartesianClosed.html#1541" class="Bound">A</a><a id="1555" class="Symbol">}</a> <a id="1557" class="Symbol">{</a><a id="1558" href="Categories.Category.CartesianClosed.html#1537" class="Bound">B</a><a id="1559" class="Symbol">}</a>
<a id="CartesianClosed._⇨_"></a><a id="1564" href="Categories.Category.CartesianClosed.html#1564" class="Function Operator">_⇨_</a> <a id="1568" class="Symbol">:</a> <a id="1570" href="Categories.Category.Core.html#559" class="Field">Obj</a> <a id="1574" class="Symbol"></a> <a id="1576" href="Categories.Category.Core.html#559" class="Field">Obj</a> <a id="1580" class="Symbol"></a> <a id="1582" href="Categories.Category.Core.html#559" class="Field">Obj</a>
<a id="1588" href="Categories.Category.CartesianClosed.html#1564" class="Function Operator">_⇨_</a> <a id="1592" class="Symbol">=</a> <a id="1594" href="Function.Base.html#1638" class="Function">flip</a> <a id="1599" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">_^_</a>
<a id="1606" class="Keyword">private</a>
<a id="1618" class="Keyword">module</a> <a id="CartesianClosed.cartesian"></a><a id="1625" href="Categories.Category.CartesianClosed.html#1625" class="Module">cartesian</a> <a id="1635" class="Symbol">=</a> <a id="1637" href="Categories.Category.Cartesian.html#727" class="Module">Cartesian</a> <a id="1647" href="Categories.Category.CartesianClosed.html#1408" class="Field">cartesian</a>
<a id="1660" class="Keyword">open</a> <a id="1665" href="Categories.Category.Cartesian.Monoidal.html#1139" class="Module">CartesianMonoidal</a> <a id="1683" href="Categories.Category.CartesianClosed.html#1408" class="Field">cartesian</a> <a id="1693" class="Keyword">using</a> <a id="1699" class="Symbol">(</a><a id="1700" href="Categories.Category.Cartesian.Monoidal.html#1852" class="Function">A×≅A</a><a id="1705" class="Symbol">)</a>
<a id="1709" class="Keyword">open</a> <a id="1714" href="Categories.Category.BinaryProducts.html#848" class="Module">BinaryProducts</a> <a id="1729" href="Categories.Category.Cartesian.html#801" class="Function">cartesian.products</a> <a id="1748" class="Keyword">using</a> <a id="1754" class="Symbol">(</a><a id="1755" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">_×_</a><a id="1758" class="Symbol">;</a> <a id="1760" href="Categories.Category.BinaryProducts.html#952" class="Field">product</a><a id="1767" class="Symbol">;</a> <a id="1769" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a><a id="1771" class="Symbol">;</a> <a id="1773" href="Categories.Object.Product.Core.html#546" class="Function">π₂</a><a id="1775" class="Symbol">;</a> <a id="1777" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">⟨_,_⟩</a><a id="1782" class="Symbol">;</a>
<a id="1788" href="Categories.Object.Product.Core.html#603" class="Function">project₁</a><a id="1796" class="Symbol">;</a> <a id="1798" href="Categories.Object.Product.Core.html#637" class="Function">project₂</a><a id="1806" class="Symbol">;</a> <a id="1808" href="Categories.Object.Product.Core.html#797" class="Function">η</a><a id="1809" class="Symbol">;</a> <a id="1811" href="Categories.Object.Product.Core.html#854" class="Function">⟨⟩-cong₂</a><a id="1819" class="Symbol">;</a> <a id="1821" href="Categories.Category.BinaryProducts.html#2975" class="Function">⟨⟩∘</a><a id="1824" class="Symbol">;</a> <a id="1826" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">_⁂_</a><a id="1829" class="Symbol">;</a> <a id="1831" href="Categories.Category.BinaryProducts.html#1963" class="Function">⟨⟩-congˡ</a><a id="1839" class="Symbol">;</a> <a id="1841" href="Categories.Category.BinaryProducts.html#1878" class="Function">⟨⟩-congʳ</a><a id="1849" class="Symbol">;</a>
<a id="1855" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a><a id="1866" class="Symbol">;</a> <a id="1868" href="Categories.Category.BinaryProducts.html#3843" class="Function">firstid</a><a id="1875" class="Symbol">;</a> <a id="1877" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a><a id="1882" class="Symbol">;</a> <a id="1884" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a><a id="1890" class="Symbol">;</a> <a id="1892" href="Categories.Category.BinaryProducts.html#3714" class="Function">first↔second</a><a id="1904" class="Symbol">;</a> <a id="1906" href="Categories.Category.BinaryProducts.html#3163" class="Function">second∘second</a><a id="1919" class="Symbol">;</a> <a id="1921" href="Categories.Category.BinaryProducts.html#2515" class="Function">⁂-cong₂</a><a id="1928" class="Symbol">;</a> <a id="1930" href="Categories.Category.BinaryProducts.html#7559" class="Function Operator">-×_</a><a id="1933" class="Symbol">)</a>
<a id="1937" class="Keyword">open</a> <a id="1942" href="Categories.Object.Terminal.html#860" class="Module">Terminal</a> <a id="1951" href="Categories.Category.Cartesian.html#777" class="Function">cartesian.terminal</a> <a id="1970" class="Keyword">using</a> <a id="1976" class="Symbol">(</a><a id="1977" href="Categories.Object.Terminal.html#905" class="Field"></a><a id="1978" class="Symbol">;</a> <a id="1980" href="Categories.Object.Terminal.html#577" class="Function">!</a><a id="1981" class="Symbol">;</a> <a id="1983" href="Categories.Object.Terminal.html#647" class="Function">!-unique₂</a><a id="1992" class="Symbol">;</a> <a id="1994" href="Categories.Object.Terminal.html#803" class="Function">-id</a><a id="1998" class="Symbol">)</a>
<a id="CartesianClosed.B^A×A"></a><a id="2003" href="Categories.Category.CartesianClosed.html#2003" class="Function">B^A×A</a> <a id="2009" class="Symbol">:</a> <a id="2011" class="Symbol"></a> <a id="2013" href="Categories.Category.CartesianClosed.html#2013" class="Bound">B</a> <a id="2015" href="Categories.Category.CartesianClosed.html#2015" class="Bound">A</a> <a id="2017" class="Symbol"></a> <a id="2019" href="Categories.Object.Product.Core.html#435" class="Record">Product</a> <a id="2027" class="Symbol">(</a><a id="2028" href="Categories.Category.CartesianClosed.html#2013" class="Bound">B</a> <a id="2030" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">^</a> <a id="2032" href="Categories.Category.CartesianClosed.html#2015" class="Bound">A</a><a id="2033" class="Symbol">)</a> <a id="2035" href="Categories.Category.CartesianClosed.html#2015" class="Bound">A</a>
<a id="2039" href="Categories.Category.CartesianClosed.html#2003" class="Function">B^A×A</a> <a id="2045" href="Categories.Category.CartesianClosed.html#2045" class="Bound">B</a> <a id="2047" href="Categories.Category.CartesianClosed.html#2047" class="Bound">A</a> <a id="2049" class="Symbol">=</a> <a id="2051" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="2063" class="Symbol">{</a><a id="2064" href="Categories.Category.CartesianClosed.html#2047" class="Bound">A</a><a id="2065" class="Symbol">}</a> <a id="2067" class="Symbol">{</a><a id="2068" href="Categories.Category.CartesianClosed.html#2045" class="Bound">B</a><a id="2069" class="Symbol">}</a>
<a id="CartesianClosed.eval"></a><a id="2074" href="Categories.Category.CartesianClosed.html#2074" class="Function">eval</a> <a id="2079" class="Symbol">:</a> <a id="2081" href="Categories.Object.Product.Core.html#510" class="Field">Product.A×B</a> <a id="2093" class="Symbol">(</a><a id="2094" href="Categories.Category.CartesianClosed.html#2003" class="Function">B^A×A</a> <a id="2100" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a> <a id="2102" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a><a id="2103" class="Symbol">)</a> <a id="2105" href="Categories.Category.Core.html#575" class="Field Operator"></a> <a id="2107" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a>
<a id="2111" href="Categories.Category.CartesianClosed.html#2074" class="Function">eval</a> <a id="2116" class="Symbol">=</a> <a id="2118" href="Categories.Object.Exponential.html#767" class="Function">exp.eval</a>
<a id="CartesianClosed.λg"></a><a id="2130" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2133" class="Symbol">:</a> <a id="2135" href="Categories.Category.CartesianClosed.html#1112" class="Generalizable">C</a> <a id="2137" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="2139" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="2141" href="Categories.Category.Core.html#575" class="Field Operator"></a> <a id="2143" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a> <a id="2145" class="Symbol"></a> <a id="2147" href="Categories.Category.CartesianClosed.html#1112" class="Generalizable">C</a> <a id="2149" href="Categories.Category.Core.html#575" class="Field Operator"></a> <a id="2151" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a> <a id="2153" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">^</a> <a id="2155" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a>
<a id="2159" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2162" href="Categories.Category.CartesianClosed.html#2162" class="Bound">f</a> <a id="2164" class="Symbol">=</a> <a id="2166" href="Categories.Object.Exponential.html#792" class="Function">exp.λg</a> <a id="2173" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="2181" href="Categories.Category.CartesianClosed.html#2162" class="Bound">f</a>
<a id="CartesianClosed.λ-cong"></a><a id="2186" href="Categories.Category.CartesianClosed.html#2186" class="Function">λ-cong</a> <a id="2193" class="Symbol">:</a> <a id="2195" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2197" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2199" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="2201" class="Symbol"></a> <a id="2203" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2206" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2208" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2210" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2213" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a>
<a id="2217" href="Categories.Category.CartesianClosed.html#2186" class="Function">λ-cong</a> <a id="2224" href="Categories.Category.CartesianClosed.html#2224" class="Bound">eq</a> <a id="2227" class="Symbol">=</a> <a id="2229" href="Categories.Object.Exponential.html#1272" class="Function">exp.λ-cong</a> <a id="2240" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="2248" href="Categories.Category.CartesianClosed.html#2224" class="Bound">eq</a>
<a id="CartesianClosed.λ-inj"></a><a id="2254" href="Categories.Category.CartesianClosed.html#2254" class="Function">λ-inj</a> <a id="2260" class="Symbol">:</a> <a id="2262" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2265" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2267" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2269" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2272" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="2274" class="Symbol"></a> <a id="2276" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2278" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2280" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a>
<a id="2284" href="Categories.Category.CartesianClosed.html#2254" class="Function">λ-inj</a> <a id="2290" class="Symbol">=</a> <a id="2292" href="Categories.Object.Exponential.html#1427" class="Function">exp.λ-inj</a> <a id="2302" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a>
<a id="CartesianClosed._×id"></a><a id="2313" href="Categories.Category.CartesianClosed.html#2313" class="Function Operator">_×id</a> <a id="2318" class="Symbol">:</a> <a id="2320" class="Symbol">(</a><a id="2321" href="Categories.Category.CartesianClosed.html#2321" class="Bound">f</a> <a id="2323" class="Symbol">:</a> <a id="2325" href="Categories.Category.CartesianClosed.html#1112" class="Generalizable">C</a> <a id="2327" href="Categories.Category.Core.html#575" class="Field Operator"></a> <a id="2329" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a> <a id="2331" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">^</a> <a id="2333" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a><a id="2334" class="Symbol">)</a> <a id="2336" class="Symbol"></a> <a id="2338" href="Categories.Category.CartesianClosed.html#1112" class="Generalizable">C</a> <a id="2340" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="2342" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="2344" href="Categories.Category.Core.html#575" class="Field Operator"></a> <a id="2346" href="Categories.Object.Product.Morphisms.html#564" class="Function Operator">[[</a> <a id="2349" href="Categories.Category.CartesianClosed.html#2003" class="Function">B^A×A</a> <a id="2355" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a> <a id="2357" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="2359" href="Categories.Object.Product.Morphisms.html#564" class="Function Operator">]]</a>
<a id="2364" href="Categories.Category.CartesianClosed.html#2364" class="Bound">f</a> <a id="2366" href="Categories.Category.CartesianClosed.html#2313" class="Function Operator">×id</a> <a id="2370" class="Symbol">=</a> <a id="2372" href="Categories.Object.Product.Morphisms.html#4276" class="Function Operator">[</a> <a id="2374" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="2382" href="Categories.Object.Product.Morphisms.html#4276" class="Function Operator"></a> <a id="2384" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="2396" href="Categories.Object.Product.Morphisms.html#4276" class="Function Operator">]</a> <a id="2398" href="Categories.Category.CartesianClosed.html#2364" class="Bound">f</a> <a id="2400" href="Categories.Object.Product.Morphisms.html#4276" class="Function Operator">×id</a>
<a id="CartesianClosed.β"></a><a id="2407" href="Categories.Category.CartesianClosed.html#2407" class="Function">β</a> <a id="2409" class="Symbol">:</a> <a id="2411" href="Categories.Category.CartesianClosed.html#2074" class="Function">eval</a> <a id="2416" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="2418" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2421" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2423" href="Categories.Category.CartesianClosed.html#2313" class="Function Operator">×id</a> <a id="2427" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2429" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a>
<a id="2433" href="Categories.Category.CartesianClosed.html#2407" class="Function">β</a> <a id="2435" class="Symbol">=</a> <a id="2437" href="Categories.Object.Exponential.html#865" class="Function">exp.β</a> <a id="2443" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a>
<a id="CartesianClosed.subst"></a><a id="2454" href="Categories.Category.CartesianClosed.html#2454" class="Function">subst</a> <a id="2460" class="Symbol">:</a> <a id="2462" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2465" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2467" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="2469" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="2471" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2473" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2476" class="Symbol">(</a><a id="2477" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2479" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="2481" class="Symbol">(</a><a id="2482" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="2484" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="2486" href="Categories.Category.Core.html#630" class="Field">id</a><a id="2488" class="Symbol">))</a>
<a id="2493" href="Categories.Category.CartesianClosed.html#2454" class="Function">subst</a> <a id="2499" class="Symbol">=</a> <a id="2501" href="Categories.Object.Exponential.html#1795" class="Function">exp.subst</a> <a id="2511" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="2519" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a>
<a id="CartesianClosed.η-exp"></a><a id="2530" href="Categories.Category.CartesianClosed.html#2530" class="Function">η-exp</a> <a id="2536" class="Symbol">:</a> <a id="2538" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2541" class="Symbol">(</a><a id="2542" href="Categories.Category.CartesianClosed.html#2074" class="Function">eval</a> <a id="2547" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="2549" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2551" href="Categories.Category.CartesianClosed.html#2313" class="Function Operator">×id</a><a id="2554" class="Symbol">)</a> <a id="2556" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2558" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a>
<a id="2562" href="Categories.Category.CartesianClosed.html#2530" class="Function">η-exp</a> <a id="2568" class="Symbol">=</a> <a id="2570" href="Categories.Object.Exponential.html#1140" class="Function">exp.η</a> <a id="2576" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a>
<a id="CartesianClosed.λ-unique"></a><a id="2587" href="Categories.Category.CartesianClosed.html#2587" class="Function">λ-unique</a> <a id="2596" class="Symbol">:</a> <a id="2598" href="Categories.Category.CartesianClosed.html#2074" class="Function">eval</a> <a id="2603" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="2605" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2607" href="Categories.Category.CartesianClosed.html#2313" class="Function Operator">×id</a> <a id="2611" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2613" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="2615" class="Symbol"></a> <a id="2617" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2619" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2621" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="2624" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a>
<a id="2628" href="Categories.Category.CartesianClosed.html#2587" class="Function">λ-unique</a> <a id="2637" class="Symbol">=</a> <a id="2639" href="Categories.Object.Exponential.html#991" class="Function">exp.λ-unique</a> <a id="2652" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a>
<a id="CartesianClosed.λ-unique₂"></a><a id="2663" href="Categories.Category.CartesianClosed.html#2663" class="Function">λ-unique₂</a> <a id="2673" class="Symbol">:</a> <a id="2675" href="Categories.Category.CartesianClosed.html#2074" class="Function">eval</a> <a id="2680" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="2682" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2684" href="Categories.Category.CartesianClosed.html#2313" class="Function Operator">×id</a> <a id="2688" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2690" href="Categories.Category.CartesianClosed.html#2074" class="Function">eval</a> <a id="2695" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="2697" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="2699" href="Categories.Category.CartesianClosed.html#2313" class="Function Operator">×id</a> <a id="2703" class="Symbol"></a> <a id="2705" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="2707" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="2709" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a>
<a id="2713" href="Categories.Category.CartesianClosed.html#2663" class="Function">λ-unique₂</a> <a id="2723" class="Symbol">=</a> <a id="2725" href="Categories.Object.Exponential.html#2507" class="Function">exp.λ-unique</a> <a id="2739" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a>
<a id="2750" class="Comment">-- the annoying detail is that B^A×A is NOT the same as B ^ A × A, but they are isomorphic.</a>
<a id="2844" class="Comment">-- make some infra so that the latter (which is more intuitive) can be used.</a>
<a id="CartesianClosed.B^A×A-iso"></a><a id="2924" href="Categories.Category.CartesianClosed.html#2924" class="Function">B^A×A-iso</a> <a id="2934" class="Symbol">:</a> <a id="2936" href="Categories.Object.Product.Core.html#510" class="Field">Product.A×B</a> <a id="2948" class="Symbol">(</a><a id="2949" href="Categories.Category.CartesianClosed.html#2003" class="Function">B^A×A</a> <a id="2955" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a> <a id="2957" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a><a id="2958" class="Symbol">)</a> <a id="2960" href="Categories.Morphism.html#1958" class="Record Operator"></a> <a id="2962" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a> <a id="2964" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">^</a> <a id="2966" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="2968" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="2970" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a>
<a id="2974" href="Categories.Category.CartesianClosed.html#2924" class="Function">B^A×A-iso</a> <a id="2984" class="Symbol">{</a><a id="2985" class="Argument">B</a> <a id="2987" class="Symbol">=</a> <a id="2989" href="Categories.Category.CartesianClosed.html#2989" class="Bound">B</a><a id="2990" class="Symbol">}</a> <a id="2992" class="Symbol">{</a><a id="2993" class="Argument">A</a> <a id="2995" class="Symbol">=</a> <a id="2997" href="Categories.Category.CartesianClosed.html#2997" class="Bound">A</a><a id="2998" class="Symbol">}</a> <a id="3000" class="Symbol">=</a> <a id="3002" class="Keyword">record</a>
<a id="3013" class="Symbol">{</a> <a id="3015" href="Categories.Morphism.html#2006" class="Field">from</a> <a id="3020" class="Symbol">=</a> <a id="3022" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3029" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3041" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a>
<a id="3053" class="Symbol">;</a> <a id="3055" href="Categories.Morphism.html#2023" class="Field">to</a> <a id="3060" class="Symbol">=</a> <a id="3062" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3069" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="3077" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a>
<a id="3093" class="Symbol">;</a> <a id="3095" href="Categories.Morphism.html#2040" class="Field">iso</a> <a id="3100" class="Symbol">=</a> <a id="3102" class="Keyword">record</a>
<a id="3115" class="Symbol">{</a> <a id="3117" href="Categories.Morphism.html#1586" class="Field">isoˡ</a> <a id="3122" class="Symbol">=</a> <a id="3124" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="3138" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3145" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="3153" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3165" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="3167" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3174" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3186" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a>
<a id="3204" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3207" href="Categories.Object.Product.Morphisms.html#2671" class="Function Operator">[</a> <a id="3209" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3221" href="Categories.Object.Product.Morphisms.html#2671" class="Function Operator">]⟨⟩∘</a> <a id="3226" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="3236" href="Categories.Object.Product.Morphisms.html#615" class="Function Operator">[</a> <a id="3238" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3250" href="Categories.Object.Product.Morphisms.html#615" class="Function Operator">]⟨</a> <a id="3253" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="3256" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="3258" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3265" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3277" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="3285" href="Categories.Object.Product.Morphisms.html#615" class="Function Operator">,</a> <a id="3287" href="Categories.Object.Product.Core.html#546" class="Function">π₂</a> <a id="3290" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="3292" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3299" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3311" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="3319" href="Categories.Object.Product.Morphisms.html#615" class="Function Operator"></a>
<a id="3331" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3334" href="Categories.Object.Product.Core.html#854" class="Function">Product.⟨⟩-cong₂</a> <a id="3351" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3363" href="Categories.Object.Product.Core.html#603" class="Function">project₁</a> <a id="3372" href="Categories.Object.Product.Core.html#637" class="Function">project₂</a> <a id="3381" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="3391" href="Categories.Object.Product.Morphisms.html#615" class="Function Operator">[</a> <a id="3393" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3405" href="Categories.Object.Product.Morphisms.html#615" class="Function Operator">]⟨</a> <a id="3408" href="Categories.Object.Product.Morphisms.html#714" class="Function Operator">[</a> <a id="3410" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3422" href="Categories.Object.Product.Morphisms.html#714" class="Function Operator">]π₁</a> <a id="3426" href="Categories.Object.Product.Morphisms.html#615" class="Function Operator">,</a> <a id="3428" href="Categories.Object.Product.Morphisms.html#776" class="Function Operator">[</a> <a id="3430" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3442" href="Categories.Object.Product.Morphisms.html#776" class="Function Operator">]π₂</a> <a id="3446" href="Categories.Object.Product.Morphisms.html#615" class="Function Operator"></a>
<a id="3458" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3461" href="Categories.Object.Product.Core.html#797" class="Function">Product.η</a> <a id="3471" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3483" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="3493" href="Categories.Category.Core.html#630" class="Field">id</a>
<a id="3506" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="3514" class="Symbol">;</a> <a id="3516" href="Categories.Morphism.html#1612" class="Field">isoʳ</a> <a id="3521" class="Symbol">=</a> <a id="3523" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="3537" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3544" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3556" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="3564" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="3566" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3573" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="3581" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a>
<a id="3603" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3606" href="Categories.Category.BinaryProducts.html#2975" class="Function">⟨⟩∘</a> <a id="3610" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="3620" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="3622" href="Categories.Object.Product.Morphisms.html#714" class="Function Operator">[</a> <a id="3624" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3636" href="Categories.Object.Product.Morphisms.html#714" class="Function Operator">]π₁</a> <a id="3640" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="3642" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3649" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="3657" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3669" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="3671" href="Categories.Object.Product.Morphisms.html#776" class="Function Operator">[</a> <a id="3673" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3685" href="Categories.Object.Product.Morphisms.html#776" class="Function Operator">]π₂</a> <a id="3689" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="3691" href="Categories.Object.Product.Core.html#2078" class="Function">repack</a> <a id="3698" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="3706" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="3718" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a>
<a id="3730" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3733" href="Categories.Object.Product.Core.html#854" class="Function">⟨⟩-cong₂</a> <a id="3742" class="Symbol">(</a><a id="3743" href="Categories.Object.Product.Core.html#603" class="Field">Product.project₁</a> <a id="3760" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a><a id="3771" class="Symbol">)</a> <a id="3773" class="Symbol">(</a><a id="3774" href="Categories.Object.Product.Core.html#637" class="Field">Product.project₂</a> <a id="3791" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a><a id="3802" class="Symbol">)</a> <a id="3804" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="3814" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="3816" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="3819" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="3821" href="Categories.Object.Product.Core.html#546" class="Function">π₂</a> <a id="3824" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a>
<a id="3836" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3839" href="Categories.Object.Product.Core.html#797" class="Function">η</a> <a id="3841" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="3851" href="Categories.Category.Core.html#630" class="Field">id</a>
<a id="3864" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="3872" class="Symbol">}</a>
<a id="3878" class="Symbol">}</a>
<a id="CartesianClosed.eval"></a><a id="3883" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="3889" class="Symbol">:</a> <a id="3891" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a> <a id="3893" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">^</a> <a id="3895" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="3897" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="3899" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="3901" href="Categories.Category.Core.html#575" class="Field Operator"></a> <a id="3903" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a>
<a id="3907" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="3913" class="Symbol">=</a> <a id="3915" href="Categories.Category.CartesianClosed.html#2074" class="Function">eval</a> <a id="3920" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="3922" href="Categories.Morphism.html#2023" class="Function">to</a>
<a id="3929" class="Keyword">where</a> <a id="3935" class="Keyword">open</a> <a id="3940" href="Categories.Morphism.html#1958" class="Module Operator">_≅_</a> <a id="3944" href="Categories.Category.CartesianClosed.html#2924" class="Function">B^A×A-iso</a>
<a id="CartesianClosed.λ-unique"></a><a id="3957" href="Categories.Category.CartesianClosed.html#3957" class="Function">λ-unique</a> <a id="3967" class="Symbol">:</a> <a id="3969" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="3975" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="3977" class="Symbol">(</a><a id="3978" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="3980" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="3982" href="Categories.Category.Core.html#630" class="Field">id</a><a id="3984" class="Symbol">)</a> <a id="3986" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="3988" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="3990" class="Symbol"></a> <a id="3992" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="3994" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="3996" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="3999" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a>
<a id="4003" href="Categories.Category.CartesianClosed.html#3957" class="Function">λ-unique</a> <a id="4013" href="Categories.Category.CartesianClosed.html#4013" class="Bound">eq</a> <a id="4016" class="Symbol">=</a> <a id="4018" href="Categories.Object.Exponential.html#991" class="Function">exp.λ-unique</a> <a id="4031" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="4039" class="Symbol">(</a><a id="4040" href="Categories.Category.Core.html#3005" class="Function"></a> <a id="4042" class="Symbol">(</a><a id="4043" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="4049" href="Categories.Object.Product.Morphisms.html#3660" class="Function Operator">[</a> <a id="4051" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="4059" href="Categories.Object.Product.Morphisms.html#3660" class="Function Operator"></a> <a id="4061" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="4069" href="Categories.Object.Product.Morphisms.html#3660" class="Function Operator"></a> <a id="4071" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="4083" href="Categories.Object.Product.Morphisms.html#3660" class="Function Operator">]repack∘×</a><a id="4092" class="Symbol">)</a> <a id="4094" href="Categories.Category.Core.html#3061" class="Function Operator"></a> <a id="4096" href="Categories.Category.CartesianClosed.html#4013" class="Bound">eq</a><a id="4098" class="Symbol">)</a>
<a id="CartesianClosed.λ-unique₂"></a><a id="4103" href="Categories.Category.CartesianClosed.html#4103" class="Function">λ-unique₂</a> <a id="4114" class="Symbol">:</a> <a id="4116" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="4122" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4124" class="Symbol">(</a><a id="4125" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="4127" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="4129" href="Categories.Category.Core.html#630" class="Field">id</a><a id="4131" class="Symbol">)</a> <a id="4133" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="4135" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="4141" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4143" class="Symbol">(</a><a id="4144" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="4146" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="4148" href="Categories.Category.Core.html#630" class="Field">id</a><a id="4150" class="Symbol">)</a> <a id="4152" class="Symbol"></a> <a id="4154" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="4156" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="4158" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a>
<a id="4162" href="Categories.Category.CartesianClosed.html#4103" class="Function">λ-unique₂</a> <a id="4173" href="Categories.Category.CartesianClosed.html#4173" class="Bound">eq</a> <a id="4176" class="Symbol">=</a> <a id="4178" class="Symbol">(</a><a id="4179" href="Categories.Category.CartesianClosed.html#3957" class="Function">λ-unique</a> <a id="4189" href="Categories.Category.CartesianClosed.html#4173" class="Bound">eq</a><a id="4191" class="Symbol">)</a> <a id="4193" href="Categories.Category.Core.html#3061" class="Function Operator"></a> <a id="4195" href="Categories.Category.Core.html#3005" class="Function"></a> <a id="4197" class="Symbol">(</a><a id="4198" href="Categories.Category.CartesianClosed.html#3957" class="Function">λ-unique</a> <a id="4208" href="Relation.Binary.Structures.html#1596" class="Function">refl</a><a id="4212" class="Symbol">)</a>
<a id="CartesianClosed.β′"></a><a id="4217" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="4220" class="Symbol">:</a> <a id="4222" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="4228" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4230" class="Symbol">(</a><a id="4231" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="4234" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="4236" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="4238" href="Categories.Category.Core.html#630" class="Field">id</a><a id="4240" class="Symbol">)</a> <a id="4242" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="4244" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a>
<a id="4248" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="4251" class="Symbol">{</a><a id="4252" class="Argument">f</a> <a id="4254" class="Symbol">=</a> <a id="4256" href="Categories.Category.CartesianClosed.html#4256" class="Bound">f</a><a id="4257" class="Symbol">}</a> <a id="4259" class="Symbol">=</a> <a id="4261" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="4271" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="4277" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4279" class="Symbol">(</a><a id="4280" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="4283" href="Categories.Category.CartesianClosed.html#4256" class="Bound">f</a> <a id="4285" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="4287" href="Categories.Category.Core.html#630" class="Field">id</a><a id="4289" class="Symbol">)</a> <a id="4291" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4294" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="4300" href="Categories.Object.Product.Morphisms.html#3660" class="Function Operator">[</a> <a id="4302" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="4310" href="Categories.Object.Product.Morphisms.html#3660" class="Function Operator"></a> <a id="4312" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a> <a id="4320" href="Categories.Object.Product.Morphisms.html#3660" class="Function Operator"></a> <a id="4322" href="Categories.Object.Exponential.html#657" class="Function">exp.product</a> <a id="4334" href="Categories.Object.Product.Morphisms.html#3660" class="Function Operator">]repack∘×</a> <a id="4344" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="4350" href="Categories.Category.CartesianClosed.html#2074" class="Function">eval</a> <a id="4355" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4357" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="4360" href="Categories.Category.CartesianClosed.html#4256" class="Bound">f</a> <a id="4362" href="Categories.Category.CartesianClosed.html#2313" class="Function Operator">×id</a> <a id="4370" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4373" href="Categories.Category.CartesianClosed.html#2407" class="Function">β</a> <a id="4375" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="4381" href="Categories.Category.CartesianClosed.html#4256" class="Bound">f</a> <a id="4401" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="CartesianClosed.η-exp"></a><a id="4406" href="Categories.Category.CartesianClosed.html#4406" class="Function">η-exp</a> <a id="4413" class="Symbol">:</a> <a id="4415" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="4418" class="Symbol">(</a><a id="4419" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="4425" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4427" class="Symbol">(</a><a id="4428" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="4430" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="4432" href="Categories.Category.Core.html#630" class="Field">id</a><a id="4434" class="Symbol">))</a> <a id="4437" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="4439" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a>
<a id="4443" href="Categories.Category.CartesianClosed.html#4406" class="Function">η-exp</a> <a id="4450" class="Symbol">=</a> <a id="4452" href="Relation.Binary.Structures.html#1622" class="Function">sym</a> <a id="4456" class="Symbol">(</a><a id="4457" href="Categories.Category.CartesianClosed.html#3957" class="Function">λ-unique</a> <a id="4467" href="Relation.Binary.Structures.html#1596" class="Function">refl</a><a id="4471" class="Symbol">)</a>
<a id="CartesianClosed.η-id"></a><a id="4476" href="Categories.Category.CartesianClosed.html#4476" class="Function">η-id</a> <a id="4482" class="Symbol">:</a> <a id="4484" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="4487" class="Symbol">(</a><a id="4488" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="4494" class="Symbol">{</a><a id="4495" class="Argument">B</a> <a id="4497" class="Symbol">=</a> <a id="4499" href="Categories.Category.CartesianClosed.html#1110" class="Generalizable">B</a><a id="4500" class="Symbol">}</a> <a id="4502" class="Symbol">{</a><a id="4503" class="Argument">A</a> <a id="4505" class="Symbol">=</a> <a id="4507" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a><a id="4508" class="Symbol">})</a> <a id="4511" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="4513" href="Categories.Category.Core.html#630" class="Field">id</a>
<a id="4518" href="Categories.Category.CartesianClosed.html#4476" class="Function">η-id</a> <a id="4524" class="Symbol">=</a> <a id="4526" href="Relation.Binary.Structures.html#1622" class="Function">sym</a> <a id="4530" class="Symbol">(</a><a id="4531" href="Categories.Category.CartesianClosed.html#3957" class="Function">λ-unique</a> <a id="4541" class="Symbol">(</a><a id="4542" href="Categories.Morphism.Reasoning.Core.html#2786" class="Function">elimʳ</a> <a id="4548" class="Symbol">(</a><a id="4549" href="Categories.Object.Product.Morphisms.html#1062" class="Function">id×id</a> <a id="4555" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a><a id="4562" class="Symbol">)))</a>
<a id="CartesianClosed.^A≅"></a><a id="4569" href="Categories.Category.CartesianClosed.html#4569" class="Function">^A≅</a> <a id="4575" class="Symbol">:</a> <a id="4577" href="Categories.Object.Terminal.html#905" class="Function"></a> <a id="4579" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">^</a> <a id="4581" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="4583" href="Categories.Morphism.html#1958" class="Record Operator"></a> <a id="4585" href="Categories.Object.Terminal.html#905" class="Function"></a>
<a id="4589" href="Categories.Category.CartesianClosed.html#4569" class="Function">^A≅</a> <a id="4595" class="Symbol">=</a> <a id="4597" class="Keyword">record</a>
<a id="4608" class="Symbol">{</a> <a id="4610" href="Categories.Morphism.html#2006" class="Field">from</a> <a id="4615" class="Symbol">=</a> <a id="4617" href="Categories.Object.Terminal.html#577" class="Function">!</a>
<a id="4623" class="Symbol">;</a> <a id="4625" href="Categories.Morphism.html#2023" class="Field">to</a> <a id="4630" class="Symbol">=</a> <a id="4632" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="4635" href="Categories.Object.Terminal.html#577" class="Function">!</a>
<a id="4641" class="Symbol">;</a> <a id="4643" href="Categories.Morphism.html#2040" class="Field">iso</a> <a id="4648" class="Symbol">=</a> <a id="4650" class="Keyword">record</a>
<a id="4663" class="Symbol">{</a> <a id="4665" href="Categories.Morphism.html#1586" class="Field">isoˡ</a> <a id="4670" class="Symbol">=</a> <a id="4672" href="Categories.Category.CartesianClosed.html#2663" class="Function">λ-unique₂</a> <a id="4682" href="Categories.Object.Terminal.html#647" class="Function">!-unique₂</a>
<a id="4698" class="Symbol">;</a> <a id="4700" href="Categories.Morphism.html#1612" class="Field">isoʳ</a> <a id="4705" class="Symbol">=</a> <a id="4707" href="Categories.Object.Terminal.html#803" class="Function">-id</a> <a id="4712" class="Symbol">_</a>
<a id="4720" class="Symbol">}</a>
<a id="4726" class="Symbol">}</a>
<a id="CartesianClosed.A^≅A"></a><a id="4731" href="Categories.Category.CartesianClosed.html#4731" class="Function">A^≅A</a> <a id="4737" class="Symbol">:</a> <a id="4739" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a> <a id="4741" href="Categories.Category.CartesianClosed.html#1513" class="Function Operator">^</a> <a id="4743" href="Categories.Object.Terminal.html#905" class="Function"></a> <a id="4745" href="Categories.Morphism.html#1958" class="Record Operator"></a> <a id="4747" href="Categories.Category.CartesianClosed.html#1108" class="Generalizable">A</a>
<a id="4751" href="Categories.Category.CartesianClosed.html#4731" class="Function">A^≅A</a> <a id="4757" class="Symbol">=</a> <a id="4759" class="Keyword">record</a>
<a id="4770" class="Symbol">{</a> <a id="4772" href="Categories.Morphism.html#2006" class="Field">from</a> <a id="4777" class="Symbol">=</a> <a id="4779" class="Keyword">let</a> <a id="4783" class="Keyword">open</a> <a id="4788" href="Categories.Morphism.html#1958" class="Module Operator">_≅_</a> <a id="4792" href="Categories.Category.Cartesian.Monoidal.html#1852" class="Function">A×≅A</a> <a id="4798" class="Keyword">in</a> <a id="4801" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="4807" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4809" href="Categories.Morphism.html#2023" class="Function">to</a>
<a id="4816" class="Symbol">;</a> <a id="4818" href="Categories.Morphism.html#2023" class="Field">to</a> <a id="4823" class="Symbol">=</a> <a id="4825" class="Keyword">let</a> <a id="4829" class="Keyword">open</a> <a id="4834" href="Categories.Morphism.html#1958" class="Module Operator">_≅_</a> <a id="4838" href="Categories.Category.Cartesian.Monoidal.html#1852" class="Function">A×≅A</a> <a id="4844" class="Keyword">in</a> <a id="4847" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="4850" href="Categories.Morphism.html#2006" class="Function">from</a>
<a id="4859" class="Symbol">;</a> <a id="4861" href="Categories.Morphism.html#2040" class="Field">iso</a> <a id="4866" class="Symbol">=</a> <a id="4868" class="Keyword">record</a>
<a id="4881" class="Symbol">{</a> <a id="4883" href="Categories.Morphism.html#1586" class="Field">isoˡ</a> <a id="4888" class="Symbol">=</a> <a id="4890" href="Categories.Category.CartesianClosed.html#4103" class="Function">λ-unique₂</a> <a id="4901" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="4903" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="4917" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="4923" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4925" class="Symbol">((</a><a id="4927" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="4930" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="4933" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4935" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="4941" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="4943" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="4945" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="4948" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="4950" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="4952" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a><a id="4953" class="Symbol">)</a> <a id="4955" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="4957" href="Categories.Category.Core.html#630" class="Field">id</a><a id="4959" class="Symbol">)</a> <a id="4970" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="4974" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="4982" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a> <a id="4994" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="5004" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="5010" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5012" class="Symbol">((</a><a id="5014" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="5017" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5020" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="5022" href="Categories.Category.Core.html#630" class="Field">id</a><a id="5024" class="Symbol">)</a> <a id="5026" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5028" class="Symbol">((</a><a id="5030" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="5036" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5038" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5040" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5043" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5045" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5047" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a><a id="5048" class="Symbol">)</a> <a id="5050" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="5052" href="Categories.Category.Core.html#630" class="Field">id</a><a id="5054" class="Symbol">))</a> <a id="5057" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5060" href="Categories.Morphism.Reasoning.Core.html#2048" class="Function">pullˡ</a> <a id="5066" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="5069" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="5079" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5082" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5084" class="Symbol">((</a><a id="5086" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="5092" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5094" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5096" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5099" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5101" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5103" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a><a id="5104" class="Symbol">)</a> <a id="5106" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="5108" href="Categories.Category.Core.html#630" class="Field">id</a><a id="5110" class="Symbol">)</a> <a id="5132" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5135" href="Categories.Category.CartesianClosed.html#5796" class="Function">helper</a> <a id="5142" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="5152" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="5158" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5160" class="Symbol">(</a><a id="5161" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5164" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="5166" href="Categories.Category.Core.html#630" class="Field">id</a><a id="5168" class="Symbol">)</a> <a id="5205" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="5213" class="Symbol">;</a> <a id="5215" href="Categories.Morphism.html#1612" class="Field">isoʳ</a> <a id="5220" class="Symbol">=</a> <a id="5222" href="Categories.Category.BinaryProducts.html#3843" class="Function">firstid</a> <a id="5230" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5232" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="5234" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="5248" class="Symbol">((</a><a id="5250" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="5256" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5258" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5260" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5263" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5265" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5267" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a><a id="5268" class="Symbol">)</a> <a id="5270" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5272" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="5275" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a><a id="5277" class="Symbol">)</a> <a id="5279" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="5281" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5301" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="5305" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a> <a id="5317" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="5327" class="Symbol">(</a><a id="5328" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="5334" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5336" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5338" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5341" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5343" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5345" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5347" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="5349" href="Categories.Category.Core.html#630" class="Field">id</a><a id="5351" class="Symbol">)</a> <a id="5353" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5355" class="Symbol">(</a><a id="5356" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="5359" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5362" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="5364" href="Categories.Category.Core.html#630" class="Field">id</a><a id="5366" class="Symbol">)</a> <a id="5380" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5383" href="Categories.Category.CartesianClosed.html#6128" class="Function">helper</a> <a id="5391" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="5399" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="5409" class="Symbol">(</a><a id="5410" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5412" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5415" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5417" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5419" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5421" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5423" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a><a id="5428" class="Symbol">)</a> <a id="5430" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5432" class="Symbol">(</a><a id="5433" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="5436" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5439" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="5441" href="Categories.Category.Core.html#630" class="Field">id</a><a id="5443" class="Symbol">)</a> <a id="5462" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5465" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="5471" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="5474" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="5484" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5486" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5489" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5491" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5493" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5495" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5497" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5537" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5540" href="Categories.Category.BinaryProducts.html#2975" class="Function">⟨⟩∘</a> <a id="5544" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="5554" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5556" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5559" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5561" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5564" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5566" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5568" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5570" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5573" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5607" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5610" href="Categories.Object.Product.Core.html#854" class="Function">⟨⟩-cong₂</a> <a id="5619" href="Categories.Category.Core.html#1096" class="Field">identityˡ</a> <a id="5629" href="Categories.Object.Terminal.html#647" class="Function">!-unique₂</a> <a id="5639" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="5649" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5651" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5654" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5656" href="Categories.Object.Product.Core.html#546" class="Function">π₂</a> <a id="5659" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5702" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5705" href="Categories.Object.Product.Core.html#797" class="Function">η</a> <a id="5707" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="5717" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5770" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="5778" class="Symbol">}</a>
<a id="5784" class="Symbol">}</a>
<a id="5790" class="Keyword">where</a> <a id="5796" href="Categories.Category.CartesianClosed.html#5796" class="Function">helper</a> <a id="5803" class="Symbol">=</a> <a id="5805" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="5823" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5826" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5828" class="Symbol">((</a><a id="5830" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="5836" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5838" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5840" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5843" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5845" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5847" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a><a id="5848" class="Symbol">)</a> <a id="5850" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="5852" href="Categories.Category.Core.html#630" class="Field">id</a><a id="5854" class="Symbol">)</a> <a id="5872" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5875" href="Categories.Object.Product.Core.html#603" class="Function">project₁</a> <a id="5884" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="5898" class="Symbol">(</a><a id="5899" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="5905" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5907" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5909" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5912" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5914" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5916" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a><a id="5917" class="Symbol">)</a> <a id="5919" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5921" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5947" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5950" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="5956" href="Categories.Category.BinaryProducts.html#2975" class="Function">⟨⟩∘</a> <a id="5960" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="5974" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="5980" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5982" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="5984" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="5987" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5989" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="5992" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="5994" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="5996" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="5998" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a> <a id="6001" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="6023" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6026" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="6034" href="Categories.Category.BinaryProducts.html#1963" class="Function">⟨⟩-congˡ</a> <a id="6043" href="Categories.Object.Terminal.html#647" class="Function">!-unique₂</a> <a id="6053" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="6067" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="6073" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6075" class="Symbol">(</a><a id="6076" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="6079" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="6081" href="Categories.Category.Core.html#630" class="Field">id</a><a id="6083" class="Symbol">)</a> <a id="6116" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="6128" href="Categories.Category.CartesianClosed.html#6128" class="Function">helper</a> <a id="6136" class="Symbol">=</a> <a id="6138" class="Keyword">let</a> <a id="6142" class="Keyword">open</a> <a id="6147" href="Categories.Morphism.html#1958" class="Module Operator">_≅_</a> <a id="6151" href="Categories.Category.Cartesian.Monoidal.html#1852" class="Function">A×≅A</a> <a id="6157" class="Keyword">in</a> <a id="6160" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="6178" class="Symbol">(</a><a id="6179" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="6185" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6187" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="6189" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="6192" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="6194" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="6196" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a><a id="6197" class="Symbol">)</a> <a id="6199" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="6201" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="6227" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6230" href="Categories.Morphism.Reasoning.Core.html#3063" class="Function">introˡ</a> <a id="6237" href="Categories.Morphism.html#1586" class="Function">isoˡ</a> <a id="6242" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="6256" class="Symbol">(</a><a id="6257" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="6259" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="6262" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="6264" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="6266" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="6268" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6270" href="Categories.Object.Product.Core.html#526" class="Function">π₁</a><a id="6272" class="Symbol">)</a> <a id="6274" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6276" class="Symbol">((</a><a id="6278" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="6284" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6286" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="6288" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="6291" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="6293" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="6295" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a><a id="6296" class="Symbol">)</a> <a id="6298" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="6300" href="Categories.Category.Core.html#630" class="Field">id</a><a id="6302" class="Symbol">)</a> <a id="6305" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6308" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="6314" href="Categories.Category.CartesianClosed.html#5796" class="Function">helper</a> <a id="6321" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="6335" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="6337" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="6340" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="6342" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="6344" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="6346" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6348" class="Symbol">(</a><a id="6349" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="6355" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6357" class="Symbol">(</a><a id="6358" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="6361" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator"></a> <a id="6363" href="Categories.Category.Core.html#630" class="Field">id</a><a id="6365" class="Symbol">))</a> <a id="6384" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6387" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="6395" href="Categories.Morphism.Reasoning.Core.html#2786" class="Function">elimʳ</a> <a id="6401" class="Symbol">(</a><a id="6402" href="Categories.Object.Product.Morphisms.html#1062" class="Function">id×id</a> <a id="6408" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a><a id="6415" class="Symbol">)</a> <a id="6417" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="6431" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="6433" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="6436" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator">,</a> <a id="6438" href="Categories.Object.Terminal.html#577" class="Function">!</a> <a id="6440" href="Categories.Category.BinaryProducts.html#1396" class="Function Operator"></a> <a id="6442" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6444" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="6480" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="6485" class="Comment">-- we use -⇨- to represent the bifunctor.</a>
<a id="6529" class="Comment">-- -^- would generate a bifunctor of type Bifunctor 𝒞 𝒞.op 𝒞 which is not very typical.</a>
<a id="CartesianClosed.-⇨-"></a><a id="6619" href="Categories.Category.CartesianClosed.html#6619" class="Function">-⇨-</a> <a id="6623" class="Symbol">:</a> <a id="6625" href="Categories.Functor.Bifunctor.html#441" class="Function">Bifunctor</a> <a id="6635" href="Categories.Category.Core.html#3132" class="Function">𝒞.op</a> <a id="6640" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a> <a id="6642" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="6646" href="Categories.Category.CartesianClosed.html#6619" class="Function">-⇨-</a> <a id="6650" class="Symbol">=</a> <a id="6652" class="Keyword">record</a>
<a id="6663" class="Symbol">{</a> <a id="6665" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="6678" class="Symbol">=</a> <a id="6680" href="Data.Product.Base.html#3109" class="Function">uncurry</a> <a id="6688" href="Categories.Category.CartesianClosed.html#1564" class="Function Operator">_⇨_</a>
<a id="6696" class="Symbol">;</a> <a id="6698" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="6711" class="Symbol">=</a> <a id="6713" class="Symbol">λ</a> <a id="6715" class="Keyword">where</a>
<a id="6727" class="Symbol">(</a><a id="6728" href="Categories.Category.CartesianClosed.html#6728" class="Bound">f</a> <a id="6730" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="6732" href="Categories.Category.CartesianClosed.html#6732" class="Bound">g</a><a id="6733" class="Symbol">)</a> <a id="6735" class="Symbol"></a> <a id="6737" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="6740" class="Symbol">(</a><a id="6741" href="Categories.Category.CartesianClosed.html#6732" class="Bound">g</a> <a id="6743" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6745" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="6751" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="6753" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="6760" href="Categories.Category.CartesianClosed.html#6728" class="Bound">f</a><a id="6761" class="Symbol">)</a>
<a id="6767" class="Symbol">;</a> <a id="6769" href="Categories.Functor.Core.html#511" class="Field">identity</a> <a id="6782" class="Symbol">=</a> <a id="6784" href="Categories.Category.CartesianClosed.html#2186" class="Function">λ-cong</a> <a id="6791" class="Symbol">(</a><a id="6792" href="Categories.Category.Core.html#1096" class="Field">identityˡ</a> <a id="6802" href="Categories.Category.Core.html#3061" class="Function Operator"></a> <a id="6804" class="Symbol">(</a><a id="6805" href="Categories.Morphism.Reasoning.Core.html#2786" class="Function">elimʳ</a> <a id="6811" class="Symbol">(</a><a id="6812" href="Categories.Object.Product.Morphisms.html#1062" class="Function">id×id</a> <a id="6818" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a><a id="6825" class="Symbol">)))</a> <a id="6829" href="Categories.Category.Core.html#3061" class="Function Operator"></a> <a id="6831" href="Categories.Category.CartesianClosed.html#4476" class="Function">η-id</a>
<a id="6841" class="Symbol">;</a> <a id="6843" href="Categories.Functor.Core.html#565" class="Field">homomorphism</a> <a id="6856" class="Symbol">=</a> <a id="6858" href="Categories.Category.CartesianClosed.html#4103" class="Function">λ-unique₂</a> <a id="6869" href="Categories.Category.CartesianClosed.html#6994" class="Function">helper</a>
<a id="6880" class="Symbol">;</a> <a id="6882" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="6895" class="Symbol">=</a> <a id="6897" class="Symbol">λ</a> <a id="6899" class="Keyword">where</a>
<a id="6911" class="Symbol">(</a><a id="6912" href="Categories.Category.CartesianClosed.html#6912" class="Bound">eq₁</a> <a id="6916" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="6918" href="Categories.Category.CartesianClosed.html#6918" class="Bound">eq₂</a><a id="6921" class="Symbol">)</a> <a id="6923" class="Symbol"></a> <a id="6925" href="Categories.Category.CartesianClosed.html#2186" class="Function">λ-cong</a> <a id="6932" class="Symbol">(</a><a id="6933" href="Categories.Category.Core.html#1438" class="Field">∘-resp-≈</a> <a id="6942" href="Categories.Category.CartesianClosed.html#6918" class="Bound">eq₂</a> <a id="6946" class="Symbol">(</a><a id="6947" href="Categories.Category.Core.html#1706" class="Function">∘-resp-≈ʳ</a> <a id="6957" class="Symbol">(</a><a id="6958" href="Categories.Category.BinaryProducts.html#2515" class="Function">⁂-cong₂</a> <a id="6966" href="Relation.Binary.Structures.html#1596" class="Function">refl</a> <a id="6971" href="Categories.Category.CartesianClosed.html#6912" class="Bound">eq₁</a><a id="6974" class="Symbol">)))</a>
<a id="6982" class="Symbol">}</a>
<a id="6988" class="Keyword">where</a> <a id="6994" href="Categories.Category.CartesianClosed.html#6994" class="Function">helper</a> <a id="7001" class="Symbol">:</a> <a id="7003" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7009" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7011" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="7017" class="Symbol">(</a><a id="7018" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="7021" class="Symbol">((</a><a id="7023" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="7025" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7027" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a><a id="7028" class="Symbol">)</a> <a id="7030" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7032" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7038" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7040" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7047" class="Symbol">(</a><a id="7048" href="Categories.Category.CartesianClosed.html#1130" class="Generalizable">h</a> <a id="7050" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7052" href="Categories.Category.CartesianClosed.html#1132" class="Generalizable">i</a><a id="7053" class="Symbol">)))</a>
<a id="7074" href="Categories.Category.Core.html#595" class="Field Operator"></a> <a id="7076" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7082" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7084" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="7090" class="Symbol">(</a><a id="7091" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="7094" class="Symbol">(</a><a id="7095" href="Categories.Category.CartesianClosed.html#1128" class="Generalizable">g</a> <a id="7097" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7099" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7105" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7107" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7114" href="Categories.Category.CartesianClosed.html#1132" class="Generalizable">i</a><a id="7115" class="Symbol">)</a> <a id="7117" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7119" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="7122" class="Symbol">(</a><a id="7123" href="Categories.Category.CartesianClosed.html#1126" class="Generalizable">f</a> <a id="7125" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7127" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7133" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7135" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7142" href="Categories.Category.CartesianClosed.html#1130" class="Generalizable">h</a><a id="7143" class="Symbol">))</a>
<a id="7156" href="Categories.Category.CartesianClosed.html#6994" class="Function">helper</a> <a id="7163" class="Symbol">{</a><a id="7164" class="Argument">g</a> <a id="7166" class="Symbol">=</a> <a id="7168" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a><a id="7169" class="Symbol">}</a> <a id="7171" class="Symbol">{</a><a id="7172" class="Argument">f</a> <a id="7174" class="Symbol">=</a> <a id="7176" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a><a id="7177" class="Symbol">}</a> <a id="7179" class="Symbol">{</a><a id="7180" class="Argument">h</a> <a id="7182" class="Symbol">=</a> <a id="7184" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a><a id="7185" class="Symbol">}</a> <a id="7187" class="Symbol">{</a><a id="7188" class="Argument">i</a> <a id="7190" class="Symbol">=</a> <a id="7192" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a><a id="7193" class="Symbol">}</a> <a id="7195" class="Symbol">=</a> <a id="7197" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="7215" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7221" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7223" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="7229" class="Symbol">(</a><a id="7230" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="7233" class="Symbol">((</a><a id="7235" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a> <a id="7237" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7239" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a><a id="7240" class="Symbol">)</a> <a id="7242" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7244" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7250" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7252" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7259" class="Symbol">(</a><a id="7260" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a> <a id="7262" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7264" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a><a id="7265" class="Symbol">)))</a> <a id="7293" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="7296" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="7299" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="7313" class="Symbol">(</a><a id="7314" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a> <a id="7316" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7318" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a><a id="7319" class="Symbol">)</a> <a id="7321" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7323" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7329" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7331" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7338" class="Symbol">(</a><a id="7339" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a> <a id="7341" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7343" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a><a id="7344" class="Symbol">)</a> <a id="7391" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="7395" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="7403" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="7409" href="Categories.Category.BinaryProducts.html#3163" class="Function">second∘second</a> <a id="7423" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="7437" class="Symbol">(</a><a id="7438" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a> <a id="7440" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7442" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a><a id="7443" class="Symbol">)</a> <a id="7445" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7447" class="Symbol">(</a><a id="7448" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7454" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7456" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7463" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a><a id="7464" class="Symbol">)</a> <a id="7466" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7468" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7475" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a> <a id="7515" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="7518" href="Categories.Morphism.Reasoning.Core.html#7259" class="Function">center</a> <a id="7525" href="Relation.Binary.Structures.html#1596" class="Function">refl</a> <a id="7530" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="7544" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a> <a id="7546" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7548" class="Symbol">(</a><a id="7549" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a> <a id="7551" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7553" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7559" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7561" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7568" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a><a id="7569" class="Symbol">)</a> <a id="7571" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7573" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7580" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a> <a id="7622" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="7626" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="7634" href="Categories.Morphism.Reasoning.Core.html#2048" class="Function">pullˡ</a> <a id="7640" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="7643" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="7657" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a> <a id="7659" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7661" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7667" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7669" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="7675" class="Symbol">(</a><a id="7676" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="7679" class="Symbol">(</a><a id="7680" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a> <a id="7682" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7684" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7690" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7692" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7699" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a><a id="7700" class="Symbol">))</a> <a id="7703" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7705" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7712" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a> <a id="7735" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="7738" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="7746" href="Categories.Morphism.Reasoning.Core.html#2347" class="Function">pushʳ</a> <a id="7752" href="Categories.Category.BinaryProducts.html#3714" class="Function">first↔second</a> <a id="7765" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="7779" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a> <a id="7781" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7783" class="Symbol">(</a><a id="7784" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7790" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7792" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7799" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a><a id="7800" class="Symbol">)</a> <a id="7802" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7804" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="7810" class="Symbol">(</a><a id="7811" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="7814" class="Symbol">(</a><a id="7815" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a> <a id="7817" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7819" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7825" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7827" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7834" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a><a id="7835" class="Symbol">))</a> <a id="7857" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="7860" href="Categories.Category.Core.html#1004" class="Field">sym-assoc</a> <a id="7870" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="7884" class="Symbol">(</a><a id="7885" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a> <a id="7887" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7889" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7895" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7897" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7904" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a><a id="7905" class="Symbol">)</a> <a id="7907" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7909" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="7915" class="Symbol">(</a><a id="7916" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="7919" class="Symbol">(</a><a id="7920" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a> <a id="7922" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7924" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7930" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7932" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="7939" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a><a id="7940" class="Symbol">))</a> <a id="7962" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="7966" href="Categories.Morphism.Reasoning.Core.html#2048" class="Function">pullˡ</a> <a id="7972" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="7975" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="7989" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="7995" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="7997" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="8003" class="Symbol">(</a><a id="8004" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="8007" class="Symbol">(</a><a id="8008" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a> <a id="8010" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8012" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="8018" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8020" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="8027" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a><a id="8028" class="Symbol">))</a> <a id="8031" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8033" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="8039" class="Symbol">(</a><a id="8040" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="8043" class="Symbol">(</a><a id="8044" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a> <a id="8046" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8048" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="8054" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8056" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="8063" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a><a id="8064" class="Symbol">))</a> <a id="8067" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8070" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="8078" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a> <a id="8090" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="8104" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="8110" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8112" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="8118" class="Symbol">(</a><a id="8119" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="8122" class="Symbol">(</a><a id="8123" href="Categories.Category.CartesianClosed.html#7168" class="Bound">g</a> <a id="8125" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8127" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="8133" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8135" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="8142" href="Categories.Category.CartesianClosed.html#7192" class="Bound">i</a><a id="8143" class="Symbol">)</a> <a id="8145" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8147" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="8150" class="Symbol">(</a><a id="8151" href="Categories.Category.CartesianClosed.html#7176" class="Bound">f</a> <a id="8153" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8155" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="8161" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="8163" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="8170" href="Categories.Category.CartesianClosed.html#7184" class="Bound">h</a><a id="8171" class="Symbol">))</a> <a id="8182" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="CartesianClosed._⇨-"></a><a id="8187" href="Categories.Category.CartesianClosed.html#8187" class="Function Operator">_⇨-</a> <a id="8191" class="Symbol">:</a> <a id="8193" href="Categories.Category.Core.html#559" class="Field">Obj</a> <a id="8197" class="Symbol"></a> <a id="8199" href="Categories.Functor.html#283" class="Function">Endofunctor</a> <a id="8211" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="8215" href="Categories.Category.CartesianClosed.html#8187" class="Function Operator">_⇨-</a> <a id="8219" class="Symbol">=</a> <a id="8221" href="Categories.Functor.Bifunctor.html#874" class="Function">appˡ</a> <a id="8226" href="Categories.Category.CartesianClosed.html#6619" class="Function">-⇨-</a>
<a id="CartesianClosed.-⇨_"></a><a id="8233" href="Categories.Category.CartesianClosed.html#8233" class="Function Operator">-⇨_</a> <a id="8237" class="Symbol">:</a> <a id="8239" href="Categories.Category.Core.html#559" class="Field">Obj</a> <a id="8243" class="Symbol"></a> <a id="8245" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="8253" href="Categories.Category.Core.html#3132" class="Function">𝒞.op</a> <a id="8258" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="8262" href="Categories.Category.CartesianClosed.html#8233" class="Function Operator">-⇨_</a> <a id="8266" class="Symbol">=</a> <a id="8268" href="Categories.Functor.Bifunctor.html#938" class="Function">appʳ</a> <a id="8273" href="Categories.Category.CartesianClosed.html#6619" class="Function">-⇨-</a>
<a id="8278" class="Comment">-- The cartesian closed structure induces a monoidal closed one:</a>
<a id="8343" class="Comment">-- 𝒞 is cartesian monoidal closed.</a>
<a id="8379" class="Keyword">module</a> <a id="CartesianMonoidalClosed"></a><a id="8386" href="Categories.Category.CartesianClosed.html#8386" class="Module">CartesianMonoidalClosed</a> <a id="8410" class="Symbol">(</a><a id="8411" href="Categories.Category.CartesianClosed.html#8411" class="Bound">cartesianClosed</a> <a id="8427" class="Symbol">:</a> <a id="8429" href="Categories.Category.CartesianClosed.html#1233" class="Record">CartesianClosed</a><a id="8444" class="Symbol">)</a> <a id="8446" class="Keyword">where</a>
<a id="8454" class="Keyword">open</a> <a id="8459" href="Categories.Category.CartesianClosed.html#1233" class="Module">CartesianClosed</a> <a id="8475" href="Categories.Category.CartesianClosed.html#8411" class="Bound">cartesianClosed</a>
<a id="8493" class="Keyword">open</a> <a id="8498" href="Categories.Category.Cartesian.Monoidal.html#1139" class="Module">CartesianMonoidal</a> <a id="8516" href="Categories.Category.CartesianClosed.html#1408" class="Field">cartesian</a> <a id="8526" class="Keyword">using</a> <a id="8532" class="Symbol">(</a><a id="8533" href="Categories.Category.Cartesian.Monoidal.html#5400" class="Function">monoidal</a><a id="8541" class="Symbol">)</a>
<a id="8545" class="Keyword">open</a> <a id="8550" href="Categories.Category.BinaryProducts.html#848" class="Module">BinaryProducts</a> <a id="8565" class="Symbol">(</a><a id="8566" href="Categories.Category.Cartesian.html#801" class="Field">Cartesian.products</a> <a id="8585" href="Categories.Category.CartesianClosed.html#1408" class="Field">cartesian</a><a id="8594" class="Symbol">)</a>
<a id="8600" class="Keyword">using</a> <a id="8606" class="Symbol">(</a><a id="8607" href="Categories.Category.BinaryProducts.html#7559" class="Function Operator">-×_</a><a id="8610" class="Symbol">;</a> <a id="8612" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a><a id="8617" class="Symbol">;</a> <a id="8619" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a><a id="8630" class="Symbol">;</a> <a id="8632" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a><a id="8638" class="Symbol">;</a> <a id="8640" href="Categories.Category.BinaryProducts.html#3714" class="Function">first↔second</a><a id="8652" class="Symbol">;</a> <a id="8654" href="Categories.Category.BinaryProducts.html#952" class="Field">product</a><a id="8661" class="Symbol">)</a>
<a id="8666" class="Keyword">private</a>
<a id="CartesianMonoidalClosed.A⇨[-×A]"></a><a id="8678" href="Categories.Category.CartesianClosed.html#8678" class="Function">A⇨[-×A]</a> <a id="8686" class="Symbol">:</a> <a id="8688" href="Categories.Category.Core.html#559" class="Field">Obj</a> <a id="8692" class="Symbol"></a> <a id="8694" href="Categories.Functor.html#283" class="Function">Endofunctor</a> <a id="8706" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="8712" href="Categories.Category.CartesianClosed.html#8678" class="Function">A⇨[-×A]</a> <a id="8720" href="Categories.Category.CartesianClosed.html#8720" class="Bound">A</a> <a id="8722" class="Symbol">=</a> <a id="8724" href="Categories.Category.CartesianClosed.html#8720" class="Bound">A</a> <a id="8726" href="Categories.Category.CartesianClosed.html#8187" class="Function Operator">⇨-</a> <a id="8729" href="Categories.Functor.html#747" class="Function Operator">∘F</a> <a id="8732" href="Categories.Category.BinaryProducts.html#7559" class="Function Operator">-×</a> <a id="8735" href="Categories.Category.CartesianClosed.html#8720" class="Bound">A</a>
<a id="8742" class="Keyword">module</a> <a id="CartesianMonoidalClosed.A⇨[-×A]"></a><a id="8749" href="Categories.Category.CartesianClosed.html#8749" class="Module">A⇨[-×A]</a> <a id="8757" class="Symbol">{</a><a id="8758" href="Categories.Category.CartesianClosed.html#8758" class="Bound">A</a><a id="8759" class="Symbol">}</a> <a id="8761" class="Symbol">=</a> <a id="8763" href="Categories.Functor.Core.html#248" class="Module">Functor</a> <a id="8771" class="Symbol">(</a><a id="8772" href="Categories.Category.CartesianClosed.html#8678" class="Function">A⇨[-×A]</a> <a id="8780" href="Categories.Category.CartesianClosed.html#8758" class="Bound">A</a><a id="8781" class="Symbol">)</a>
<a id="CartesianMonoidalClosed.[A⇨-]×A"></a><a id="8788" href="Categories.Category.CartesianClosed.html#8788" class="Function">[A⇨-]×A</a> <a id="8796" class="Symbol">:</a> <a id="8798" href="Categories.Category.Core.html#559" class="Field">Obj</a> <a id="8802" class="Symbol"></a> <a id="8804" href="Categories.Functor.html#283" class="Function">Endofunctor</a> <a id="8816" href="Categories.Category.CartesianClosed.html#121" class="Bound">𝒞</a>
<a id="8822" href="Categories.Category.CartesianClosed.html#8788" class="Function">[A⇨-]×A</a> <a id="8830" href="Categories.Category.CartesianClosed.html#8830" class="Bound">A</a> <a id="8832" class="Symbol">=</a> <a id="8834" href="Categories.Category.BinaryProducts.html#7559" class="Function Operator">-×</a> <a id="8837" href="Categories.Category.CartesianClosed.html#8830" class="Bound">A</a> <a id="8839" href="Categories.Functor.html#747" class="Function Operator">∘F</a> <a id="8842" href="Categories.Category.CartesianClosed.html#8830" class="Bound">A</a> <a id="8844" href="Categories.Category.CartesianClosed.html#8187" class="Function Operator">⇨-</a>
<a id="8852" class="Keyword">module</a> <a id="CartesianMonoidalClosed.[A⇨-]×A"></a><a id="8859" href="Categories.Category.CartesianClosed.html#8859" class="Module">[A⇨-]×A</a> <a id="8867" class="Symbol">{</a><a id="8868" href="Categories.Category.CartesianClosed.html#8868" class="Bound">A</a><a id="8869" class="Symbol">}</a> <a id="8871" class="Symbol">=</a> <a id="8873" href="Categories.Functor.Core.html#248" class="Module">Functor</a> <a id="8881" class="Symbol">(</a><a id="8882" href="Categories.Category.CartesianClosed.html#8788" class="Function">[A⇨-]×A</a> <a id="8890" href="Categories.Category.CartesianClosed.html#8868" class="Bound">A</a><a id="8891" class="Symbol">)</a>
<a id="CartesianMonoidalClosed.closedMonoidal"></a><a id="8896" href="Categories.Category.CartesianClosed.html#8896" class="Function">closedMonoidal</a> <a id="8911" class="Symbol">:</a> <a id="8913" href="Categories.Category.Monoidal.Closed.html#1618" class="Record">Closed</a> <a id="8920" href="Categories.Category.Cartesian.Monoidal.html#5400" class="Function">monoidal</a>
<a id="8931" href="Categories.Category.CartesianClosed.html#8896" class="Function">closedMonoidal</a> <a id="8946" class="Symbol">=</a> <a id="8948" class="Keyword">record</a>
<a id="8959" class="Symbol">{</a> <a id="8961" href="Categories.Category.Monoidal.Closed.html#1691" class="Field">[-,-]</a> <a id="8969" class="Symbol">=</a> <a id="8971" href="Categories.Category.CartesianClosed.html#6619" class="Function">-⇨-</a>
<a id="8979" class="Symbol">;</a> <a id="8981" href="Categories.Category.Monoidal.Closed.html#1724" class="Field">adjoint</a> <a id="8989" class="Symbol">=</a> <a id="8991" class="Symbol">λ</a> <a id="8993" class="Symbol">{</a><a id="8994" href="Categories.Category.CartesianClosed.html#8994" class="Bound">A</a><a id="8995" class="Symbol">}</a> <a id="8997" class="Symbol"></a> <a id="8999" class="Keyword">record</a>
<a id="9012" class="Symbol">{</a> <a id="9014" href="Categories.Adjoint.html#1473" class="Field">unit</a> <a id="9021" class="Symbol">=</a> <a id="9023" href="Categories.NaturalTransformation.Core.html#1750" class="Function">ntHelper</a> <a id="9032" class="Keyword">record</a>
<a id="9047" class="Symbol">{</a> <a id="9049" href="Categories.NaturalTransformation.Core.html#1637" class="Field">η</a> <a id="9057" class="Symbol">=</a> <a id="9059" class="Symbol">λ</a> <a id="9061" href="Categories.Category.CartesianClosed.html#9061" class="Bound">_</a> <a id="9063" class="Symbol"></a> <a id="9065" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="9068" href="Categories.Category.Core.html#630" class="Field">id</a>
<a id="9079" class="Symbol">;</a> <a id="9081" href="Categories.NaturalTransformation.Core.html#1681" class="Field">commute</a> <a id="9089" class="Symbol">=</a> <a id="9091" class="Symbol">λ</a> <a id="9093" href="Categories.Category.CartesianClosed.html#9093" class="Bound">f</a> <a id="9095" class="Symbol"></a> <a id="9097" href="Categories.Category.CartesianClosed.html#4103" class="Function">λ-unique₂</a> <a id="9108" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="9110" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="9126" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="9132" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9134" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9140" class="Symbol">(</a><a id="9141" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="9144" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="9147" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9149" href="Categories.Category.CartesianClosed.html#9093" class="Bound">f</a><a id="9150" class="Symbol">)</a> <a id="9172" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="9176" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="9184" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a> <a id="9196" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="9208" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="9214" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9216" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9222" class="Symbol">(</a><a id="9223" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="9226" href="Categories.Category.Core.html#630" class="Field">id</a><a id="9228" class="Symbol">)</a> <a id="9230" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9232" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9238" href="Categories.Category.CartesianClosed.html#9093" class="Bound">f</a> <a id="9254" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9257" href="Categories.Morphism.Reasoning.Core.html#6851" class="Function">cancelˡ</a> <a id="9265" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="9268" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="9280" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9286" href="Categories.Category.CartesianClosed.html#9093" class="Bound">f</a> <a id="9326" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="9330" href="Categories.Morphism.Reasoning.Core.html#6665" class="Function">cancelʳ</a> <a id="9338" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="9341" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="9353" class="Symbol">(</a><a id="9354" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9360" href="Categories.Category.CartesianClosed.html#9093" class="Bound">f</a> <a id="9362" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9364" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a><a id="9369" class="Symbol">)</a> <a id="9372" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9374" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9380" class="Symbol">(</a><a id="9381" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="9384" href="Categories.Category.Core.html#630" class="Field">id</a><a id="9386" class="Symbol">)</a> <a id="9399" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="9403" href="Categories.Category.Core.html#1706" class="Function">∘-resp-≈ʳ</a> <a id="9413" class="Symbol">(</a><a id="9414" href="Categories.Morphism.Reasoning.Core.html#2786" class="Function">elimʳ</a> <a id="9420" class="Symbol">(</a><a id="9421" href="Categories.Object.Product.Morphisms.html#1062" class="Function">id×id</a> <a id="9427" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a><a id="9434" class="Symbol">))</a> <a id="9437" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="9445" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="9457" class="Symbol">(</a><a id="9458" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9464" href="Categories.Category.CartesianClosed.html#9093" class="Bound">f</a> <a id="9466" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9468" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="9474" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9476" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9482" href="Categories.Category.Core.html#630" class="Field">id</a><a id="9484" class="Symbol">)</a> <a id="9487" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9489" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9495" class="Symbol">(</a><a id="9496" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="9499" href="Categories.Category.Core.html#630" class="Field">id</a><a id="9501" class="Symbol">)</a> <a id="9503" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="9507" href="Categories.Morphism.Reasoning.Core.html#2048" class="Function">pullˡ</a> <a id="9513" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="9516" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="9528" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="9534" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9536" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9542" class="Symbol">(</a><a id="9543" href="Categories.Functor.Core.html#455" class="Function">A⇨[-×A].F₁</a> <a id="9554" href="Categories.Category.CartesianClosed.html#9093" class="Bound">f</a><a id="9555" class="Symbol">)</a> <a id="9557" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9559" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9565" class="Symbol">(</a><a id="9566" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="9569" href="Categories.Category.Core.html#630" class="Field">id</a><a id="9571" class="Symbol">)</a> <a id="9574" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9577" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="9585" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a> <a id="9597" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="9609" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="9615" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9617" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9623" class="Symbol">(</a><a id="9624" href="Categories.Functor.Core.html#455" class="Function">A⇨[-×A].F₁</a> <a id="9635" href="Categories.Category.CartesianClosed.html#9093" class="Bound">f</a> <a id="9637" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9639" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="9642" href="Categories.Category.Core.html#630" class="Field">id</a><a id="9644" class="Symbol">)</a> <a id="9655" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="9665" class="Symbol">}</a>
<a id="9673" class="Symbol">;</a> <a id="9675" href="Categories.Adjoint.html#1521" class="Field">counit</a> <a id="9682" class="Symbol">=</a> <a id="9684" href="Categories.NaturalTransformation.Core.html#1750" class="Function">ntHelper</a> <a id="9693" class="Keyword">record</a>
<a id="9708" class="Symbol">{</a> <a id="9710" href="Categories.NaturalTransformation.Core.html#1637" class="Field">η</a> <a id="9718" class="Symbol">=</a> <a id="9720" class="Symbol">λ</a> <a id="9722" href="Categories.Category.CartesianClosed.html#9722" class="Bound">_</a> <a id="9724" class="Symbol"></a> <a id="9726" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a>
<a id="9740" class="Symbol">;</a> <a id="9742" href="Categories.NaturalTransformation.Core.html#1681" class="Field">commute</a> <a id="9750" class="Symbol">=</a> <a id="9752" class="Symbol">λ</a> <a id="9754" href="Categories.Category.CartesianClosed.html#9754" class="Bound">f</a> <a id="9756" class="Symbol"></a> <a id="9758" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="9774" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="9780" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9782" href="Categories.Functor.Core.html#455" class="Function">[A⇨-]×A.F₁</a> <a id="9793" href="Categories.Category.CartesianClosed.html#9754" class="Bound">f</a> <a id="9795" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9798" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="9801" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="9813" href="Categories.Category.CartesianClosed.html#9754" class="Bound">f</a> <a id="9815" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9817" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="9823" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9825" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9831" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="9834" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9837" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="9845" href="Categories.Morphism.Reasoning.Core.html#2786" class="Function">elimʳ</a> <a id="9851" class="Symbol">(</a><a id="9852" href="Categories.Object.Product.Morphisms.html#1062" class="Function">id×id</a> <a id="9858" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a><a id="9865" class="Symbol">)</a> <a id="9867" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="9879" href="Categories.Category.CartesianClosed.html#9754" class="Bound">f</a> <a id="9881" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9883" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="9900" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="9910" class="Symbol">}</a>
<a id="9918" class="Symbol">;</a> <a id="9920" href="Categories.Adjoint.html#1669" class="Field">zig</a> <a id="9927" class="Symbol">=</a> <a id="9929" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a>
<a id="9938" class="Symbol">;</a> <a id="9940" href="Categories.Adjoint.html#1742" class="Field">zag</a> <a id="9947" class="Symbol">=</a> <a id="9949" href="Categories.Category.CartesianClosed.html#4103" class="Function">λ-unique₂</a> <a id="9960" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="9962" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="9976" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="9982" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="9984" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="9990" class="Symbol">(</a><a id="9991" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="9994" class="Symbol">(</a><a id="9995" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10001" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10003" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10009" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10011" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10018" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10020" class="Symbol">)</a> <a id="10022" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10024" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10027" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10029" class="Symbol">)</a>
<a id="10071" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="10075" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="10083" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a> <a id="10095" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="10105" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10111" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10113" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10119" class="Symbol">(</a><a id="10120" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10123" class="Symbol">(</a><a id="10124" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10130" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10132" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10138" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10140" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10147" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10149" class="Symbol">))</a> <a id="10152" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10154" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10160" class="Symbol">(</a><a id="10161" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10164" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10166" class="Symbol">)</a>
<a id="10208" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10211" href="Categories.Morphism.Reasoning.Core.html#2048" class="Function">pullˡ</a> <a id="10217" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="10220" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="10230" class="Symbol">(</a><a id="10231" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10237" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10239" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10245" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10247" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10254" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10256" class="Symbol">)</a> <a id="10258" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10260" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10266" class="Symbol">(</a><a id="10267" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10270" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10272" class="Symbol">)</a>
<a id="10314" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10317" href="Categories.Category.Core.html#1706" class="Function">∘-resp-≈ʳ</a> <a id="10327" class="Symbol">(</a><a id="10328" href="Categories.Morphism.Reasoning.Core.html#2786" class="Function">elimʳ</a> <a id="10334" class="Symbol">(</a><a id="10335" href="Categories.Object.Product.Morphisms.html#1062" class="Function">id×id</a> <a id="10341" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a><a id="10348" class="Symbol">))</a> <a id="10351" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="10359" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="10369" class="Symbol">(</a><a id="10370" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10376" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10378" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a><a id="10383" class="Symbol">)</a> <a id="10385" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10387" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10393" class="Symbol">(</a><a id="10394" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10397" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10399" class="Symbol">)</a> <a id="10401" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10404" href="Categories.Morphism.Reasoning.Core.html#6665" class="Function">cancelʳ</a> <a id="10412" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="10415" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="10425" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10457" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="10461" href="Categories.Morphism.Reasoning.Core.html#2786" class="Function">elimʳ</a> <a id="10467" class="Symbol">(</a><a id="10468" href="Categories.Object.Product.Morphisms.html#1062" class="Function">id×id</a> <a id="10474" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a><a id="10481" class="Symbol">)</a> <a id="10483" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="10493" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10499" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10501" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10507" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="10525" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="10533" class="Symbol">}</a>
<a id="10539" class="Symbol">;</a> <a id="10541" href="Categories.Category.Monoidal.Closed.html#1760" class="Field">mate</a> <a id="10549" class="Symbol">=</a> <a id="10551" class="Symbol">λ</a> <a id="10553" class="Symbol">{</a><a id="10554" href="Categories.Category.CartesianClosed.html#10554" class="Bound">X</a> <a id="10556" href="Categories.Category.CartesianClosed.html#10556" class="Bound">Y</a><a id="10557" class="Symbol">}</a> <a id="10559" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a> <a id="10561" class="Symbol"></a> <a id="10563" class="Keyword">record</a>
<a id="10576" class="Symbol">{</a> <a id="10578" href="Categories.Adjoint.Mate.html#1486" class="Field">commute₁</a> <a id="10587" class="Symbol">=</a> <a id="10589" href="Categories.Category.CartesianClosed.html#4103" class="Function">λ-unique₂</a> <a id="10600" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="10602" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="10616" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10622" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10624" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10630" class="Symbol">(</a><a id="10631" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10634" class="Symbol">(</a><a id="10635" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10642" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a> <a id="10644" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10646" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10652" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10654" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10661" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10663" class="Symbol">)</a> <a id="10665" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10667" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10670" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10672" class="Symbol">)</a> <a id="10682" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="10686" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="10694" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a> <a id="10706" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="10716" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10722" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10724" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10730" class="Symbol">(</a><a id="10731" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10734" class="Symbol">(</a><a id="10735" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10742" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a> <a id="10744" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10746" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10752" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10754" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10761" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10763" class="Symbol">))</a> <a id="10766" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10768" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10774" class="Symbol">(</a><a id="10775" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10778" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10780" class="Symbol">)</a> <a id="10782" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10785" href="Categories.Morphism.Reasoning.Core.html#2048" class="Function">pullˡ</a> <a id="10791" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="10794" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="10804" class="Symbol">(</a><a id="10805" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10812" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a> <a id="10814" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10816" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="10822" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10824" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10831" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10833" class="Symbol">)</a> <a id="10835" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10837" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10843" class="Symbol">(</a><a id="10844" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10847" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10849" class="Symbol">)</a> <a id="10870" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10873" href="Categories.Category.Core.html#1706" class="Function">∘-resp-≈ʳ</a> <a id="10883" class="Symbol">(</a><a id="10884" href="Categories.Morphism.Reasoning.Core.html#2786" class="Function">elimʳ</a> <a id="10890" class="Symbol">(</a><a id="10891" href="Categories.Object.Product.Morphisms.html#1062" class="Function">id×id</a> <a id="10897" href="Categories.Category.BinaryProducts.html#952" class="Function">product</a><a id="10904" class="Symbol">))</a> <a id="10907" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="10915" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="10925" class="Symbol">(</a><a id="10926" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="10933" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a> <a id="10935" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10937" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a><a id="10942" class="Symbol">)</a> <a id="10944" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="10946" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="10952" class="Symbol">(</a><a id="10953" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="10956" href="Categories.Category.Core.html#630" class="Field">id</a><a id="10958" class="Symbol">)</a> <a id="10991" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10994" href="Categories.Morphism.Reasoning.Core.html#6665" class="Function">cancelʳ</a> <a id="11002" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="11005" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="11015" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="11022" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a> <a id="11081" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="11085" href="Categories.Morphism.Reasoning.Core.html#6851" class="Function">cancelˡ</a> <a id="11093" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="11096" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="11106" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11112" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11114" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="11120" class="Symbol">(</a><a id="11121" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="11124" href="Categories.Category.Core.html#630" class="Field">id</a><a id="11126" class="Symbol">)</a> <a id="11128" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11130" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="11137" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a> <a id="11172" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11175" href="Categories.Morphism.Reasoning.Core.html#2347" class="Function">pushʳ</a> <a id="11181" href="Categories.Category.BinaryProducts.html#3714" class="Function">first↔second</a> <a id="11194" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="11204" class="Symbol">(</a><a id="11205" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11211" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11213" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="11220" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a><a id="11221" class="Symbol">)</a> <a id="11223" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11225" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="11231" class="Symbol">(</a><a id="11232" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="11235" href="Categories.Category.Core.html#630" class="Field">id</a><a id="11237" class="Symbol">)</a> <a id="11270" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="11274" href="Categories.Category.Core.html#1096" class="Field">identityˡ</a> <a id="11284" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="11292" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="11302" class="Symbol">(</a><a id="11303" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="11306" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11308" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11314" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11316" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="11323" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a><a id="11324" class="Symbol">)</a> <a id="11326" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11328" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="11334" class="Symbol">(</a><a id="11335" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="11338" href="Categories.Category.Core.html#630" class="Field">id</a><a id="11340" class="Symbol">)</a> <a id="11368" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="11372" href="Categories.Morphism.Reasoning.Core.html#2048" class="Function">pullˡ</a> <a id="11378" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="11381" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="11391" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11397" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11399" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="11405" class="Symbol">(</a><a id="11406" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="11409" class="Symbol">(</a><a id="11410" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="11413" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11415" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11421" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11423" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="11430" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a><a id="11431" class="Symbol">))</a> <a id="11434" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11436" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="11442" class="Symbol">(</a><a id="11443" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="11446" href="Categories.Category.Core.html#630" class="Field">id</a><a id="11448" class="Symbol">)</a> <a id="11457" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11460" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="11468" href="Categories.Category.BinaryProducts.html#3041" class="Function">first∘first</a> <a id="11480" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="11490" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11496" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11498" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="11504" class="Symbol">(</a><a id="11505" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="11508" class="Symbol">(</a><a id="11509" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="11512" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11514" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11520" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11522" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="11529" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a><a id="11530" class="Symbol">)</a> <a id="11532" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11534" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="11537" href="Categories.Category.Core.html#630" class="Field">id</a><a id="11539" class="Symbol">)</a> <a id="11556" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="11564" class="Symbol">;</a> <a id="11566" href="Categories.Adjoint.Mate.html#1548" class="Field">commute₂</a> <a id="11575" class="Symbol">=</a> <a id="11577" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="11591" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11597" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11599" href="Categories.Category.BinaryProducts.html#2249" class="Function">first</a> <a id="11605" class="Symbol">(</a><a id="11606" href="Categories.Category.CartesianClosed.html#2130" class="Function">λg</a> <a id="11609" class="Symbol">(</a><a id="11610" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="11613" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11615" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11621" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11623" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="11630" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a><a id="11631" class="Symbol">))</a> <a id="11634" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11637" href="Categories.Category.CartesianClosed.html#4217" class="Function">β′</a> <a id="11640" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="11650" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="11653" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11655" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11661" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11663" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="11670" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a> <a id="11693" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11696" href="Categories.Category.Core.html#1096" class="Field">identityˡ</a> <a id="11706" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="11716" href="Categories.Category.CartesianClosed.html#3883" class="Function">eval</a> <a id="11722" href="Categories.Category.Core.html#656" class="Field Operator"></a> <a id="11724" href="Categories.Category.BinaryProducts.html#2301" class="Function">second</a> <a id="11731" href="Categories.Category.CartesianClosed.html#10559" class="Bound">f</a> <a id="11759" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="11767" class="Symbol">}</a>
<a id="11773" class="Symbol">}</a>
<a id="11778" class="Keyword">open</a> <a id="11783" href="Categories.Category.Monoidal.Closed.html#1618" class="Module">Closed</a> <a id="11790" href="Categories.Category.CartesianClosed.html#8896" class="Function">closedMonoidal</a> <a id="11805" class="Keyword">public</a>
</pre></body></html>