mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
73 lines
No EOL
46 KiB
HTML
73 lines
No EOL
46 KiB
HTML
<!DOCTYPE HTML>
|
||
<html><head><meta charset="utf-8"><title>Categories.Category.Distributive</title><link rel="stylesheet" href="Agda.css"></head><body><pre class="Agda"><a id="1" class="Symbol">{-#</a> <a id="5" class="Keyword">OPTIONS</a> <a id="13" class="Pragma">--without-K</a> <a id="25" class="Pragma">--safe</a> <a id="32" class="Symbol">#-}</a>
|
||
|
||
<a id="37" class="Keyword">open</a> <a id="42" class="Keyword">import</a> <a id="49" href="Level.html" class="Module">Level</a> <a id="55" class="Keyword">using</a> <a id="61" class="Symbol">(</a><a id="62" href="Level.html#602" class="Function">levelOfTerm</a><a id="73" class="Symbol">)</a>
|
||
<a id="75" class="Keyword">open</a> <a id="80" class="Keyword">import</a> <a id="87" href="Categories.Category.Core.html" class="Module">Categories.Category.Core</a> <a id="112" class="Keyword">using</a> <a id="118" class="Symbol">(</a><a id="119" href="Categories.Category.Core.html#442" class="Record">Category</a><a id="127" class="Symbol">)</a>
|
||
<a id="129" class="Keyword">open</a> <a id="134" class="Keyword">import</a> <a id="141" href="Categories.Category.Cartesian.html" class="Module">Categories.Category.Cartesian</a> <a id="171" class="Keyword">using</a> <a id="177" class="Symbol">(</a><a id="178" href="Categories.Category.Cartesian.html#727" class="Record">Cartesian</a><a id="187" class="Symbol">)</a>
|
||
<a id="189" class="Keyword">open</a> <a id="194" class="Keyword">import</a> <a id="201" href="Categories.Category.BinaryProducts.html" class="Module">Categories.Category.BinaryProducts</a> <a id="236" class="Keyword">using</a> <a id="242" class="Symbol">(</a><a id="243" href="Categories.Category.BinaryProducts.html#848" class="Record">BinaryProducts</a><a id="257" class="Symbol">)</a>
|
||
<a id="259" class="Keyword">open</a> <a id="264" class="Keyword">import</a> <a id="271" href="Categories.Category.Cocartesian.html" class="Module">Categories.Category.Cocartesian</a> <a id="303" class="Keyword">using</a> <a id="309" class="Symbol">(</a><a id="310" href="Categories.Category.Cocartesian.html#3385" class="Record">Cocartesian</a><a id="321" class="Symbol">)</a>
|
||
<a id="323" class="Keyword">import</a> <a id="330" href="Categories.Morphism.html" class="Module">Categories.Morphism</a> <a id="350" class="Symbol">as</a> <a id="353" class="Module">M</a>
|
||
<a id="355" class="Keyword">import</a> <a id="362" href="Categories.Morphism.Reasoning.html" class="Module">Categories.Morphism.Reasoning</a> <a id="392" class="Symbol">as</a> <a id="395" class="Module">MR</a>
|
||
|
||
<a id="399" class="Comment">-- A distributive category is a cartesian, cocartesian category</a>
|
||
<a id="463" class="Comment">-- where the canonical distributivity morphism is an iso</a>
|
||
<a id="520" class="Comment">-- https://ncatlab.org/nlab/show/distributive+category</a>
|
||
|
||
<a id="576" class="Keyword">module</a> <a id="583" href="Categories.Category.Distributive.html" class="Module">Categories.Category.Distributive</a> <a id="616" class="Symbol">{</a><a id="617" href="Categories.Category.Distributive.html#617" class="Bound">o</a> <a id="619" href="Categories.Category.Distributive.html#619" class="Bound">ℓ</a> <a id="621" href="Categories.Category.Distributive.html#621" class="Bound">e</a><a id="622" class="Symbol">}</a> <a id="624" class="Symbol">(</a><a id="625" href="Categories.Category.Distributive.html#625" class="Bound">𝒞</a> <a id="627" class="Symbol">:</a> <a id="629" href="Categories.Category.Core.html#442" class="Record">Category</a> <a id="638" href="Categories.Category.Distributive.html#617" class="Bound">o</a> <a id="640" href="Categories.Category.Distributive.html#619" class="Bound">ℓ</a> <a id="642" href="Categories.Category.Distributive.html#621" class="Bound">e</a><a id="643" class="Symbol">)</a> <a id="645" class="Keyword">where</a>
|
||
<a id="651" class="Keyword">open</a> <a id="656" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="665" href="Categories.Category.Distributive.html#625" class="Bound">𝒞</a>
|
||
<a id="667" class="Keyword">open</a> <a id="672" href="Categories.Morphism.html" class="Module">M</a> <a id="674" href="Categories.Category.Distributive.html#625" class="Bound">𝒞</a>
|
||
<a id="676" class="Keyword">open</a> <a id="681" href="Categories.Morphism.Reasoning.html" class="Module">MR</a> <a id="684" href="Categories.Category.Distributive.html#625" class="Bound">𝒞</a>
|
||
<a id="686" class="Keyword">open</a> <a id="691" href="Categories.Category.Core.html#2462" class="Module">HomReasoning</a>
|
||
|
||
<a id="705" class="Keyword">record</a> <a id="Distributive"></a><a id="712" href="Categories.Category.Distributive.html#712" class="Record">Distributive</a> <a id="725" class="Symbol">:</a> <a id="727" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="731" class="Symbol">(</a><a id="732" href="Level.html#602" class="Function">levelOfTerm</a> <a id="744" href="Categories.Category.Distributive.html#625" class="Bound">𝒞</a><a id="745" class="Symbol">)</a> <a id="747" class="Keyword">where</a>
|
||
<a id="755" class="Keyword">field</a>
|
||
<a id="Distributive.cartesian"></a><a id="765" href="Categories.Category.Distributive.html#765" class="Field">cartesian</a> <a id="775" class="Symbol">:</a> <a id="777" href="Categories.Category.Cartesian.html#727" class="Record">Cartesian</a> <a id="787" href="Categories.Category.Distributive.html#625" class="Bound">𝒞</a>
|
||
<a id="Distributive.cocartesian"></a><a id="793" href="Categories.Category.Distributive.html#793" class="Field">cocartesian</a> <a id="805" class="Symbol">:</a> <a id="807" href="Categories.Category.Cocartesian.html#3385" class="Record">Cocartesian</a> <a id="819" href="Categories.Category.Distributive.html#625" class="Bound">𝒞</a>
|
||
|
||
<a id="824" class="Keyword">open</a> <a id="829" href="Categories.Category.Cartesian.html#727" class="Module">Cartesian</a> <a id="839" href="Categories.Category.Distributive.html#765" class="Field">cartesian</a> <a id="849" class="Keyword">using</a> <a id="855" class="Symbol">(</a><a id="856" href="Categories.Category.Cartesian.html#801" class="Field">products</a><a id="864" class="Symbol">)</a>
|
||
<a id="868" class="Keyword">open</a> <a id="873" href="Categories.Category.BinaryProducts.html#848" class="Module">BinaryProducts</a> <a id="888" href="Categories.Category.Cartesian.html#801" class="Function">products</a>
|
||
<a id="899" class="Keyword">open</a> <a id="904" href="Categories.Category.Cocartesian.html#3385" class="Module">Cocartesian</a> <a id="916" href="Categories.Category.Distributive.html#793" class="Field">cocartesian</a>
|
||
|
||
<a id="Distributive.distributeˡ"></a><a id="931" href="Categories.Category.Distributive.html#931" class="Function">distributeˡ</a> <a id="943" class="Symbol">:</a> <a id="945" class="Symbol">∀</a> <a id="947" class="Symbol">{</a><a id="948" href="Categories.Category.Distributive.html#948" class="Bound">A</a> <a id="950" href="Categories.Category.Distributive.html#950" class="Bound">B</a> <a id="952" href="Categories.Category.Distributive.html#952" class="Bound">C</a> <a id="954" class="Symbol">:</a> <a id="956" href="Categories.Category.Core.html#559" class="Field">Obj</a><a id="959" class="Symbol">}</a> <a id="961" class="Symbol">→</a> <a id="963" href="Categories.Category.Distributive.html#948" class="Bound">A</a> <a id="965" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="967" href="Categories.Category.Distributive.html#950" class="Bound">B</a> <a id="969" href="Categories.Category.Cocartesian.html#1549" class="Function Operator">+</a> <a id="971" href="Categories.Category.Distributive.html#948" class="Bound">A</a> <a id="973" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="975" href="Categories.Category.Distributive.html#952" class="Bound">C</a> <a id="977" href="Categories.Category.Core.html#575" class="Field Operator">⇒</a> <a id="979" href="Categories.Category.Distributive.html#948" class="Bound">A</a> <a id="981" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="983" class="Symbol">(</a><a id="984" href="Categories.Category.Distributive.html#950" class="Bound">B</a> <a id="986" href="Categories.Category.Cocartesian.html#1549" class="Function Operator">+</a> <a id="988" href="Categories.Category.Distributive.html#952" class="Bound">C</a><a id="989" class="Symbol">)</a>
|
||
<a id="993" href="Categories.Category.Distributive.html#931" class="Function">distributeˡ</a> <a id="1005" class="Symbol">=</a> <a id="1007" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="1009" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1012" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1014" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a> <a id="1017" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="1019" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1022" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1024" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a> <a id="1027" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a>
|
||
|
||
<a id="1032" class="Keyword">field</a>
|
||
<a id="Distributive.isIsoˡ"></a><a id="1042" href="Categories.Category.Distributive.html#1042" class="Field">isIsoˡ</a> <a id="1049" class="Symbol">:</a> <a id="1051" class="Symbol">∀</a> <a id="1053" class="Symbol">{</a><a id="1054" href="Categories.Category.Distributive.html#1054" class="Bound">A</a> <a id="1056" href="Categories.Category.Distributive.html#1056" class="Bound">B</a> <a id="1058" href="Categories.Category.Distributive.html#1058" class="Bound">C</a> <a id="1060" class="Symbol">:</a> <a id="1062" href="Categories.Category.Core.html#559" class="Field">Obj</a><a id="1065" class="Symbol">}</a> <a id="1067" class="Symbol">→</a> <a id="1069" href="Categories.Morphism.html#1826" class="Record">IsIso</a> <a id="1075" class="Symbol">(</a><a id="1076" href="Categories.Category.Distributive.html#931" class="Function">distributeˡ</a> <a id="1088" class="Symbol">{</a><a id="1089" href="Categories.Category.Distributive.html#1054" class="Bound">A</a><a id="1090" class="Symbol">}</a> <a id="1092" class="Symbol">{</a><a id="1093" href="Categories.Category.Distributive.html#1056" class="Bound">B</a><a id="1094" class="Symbol">}</a> <a id="1096" class="Symbol">{</a><a id="1097" href="Categories.Category.Distributive.html#1058" class="Bound">C</a><a id="1098" class="Symbol">})</a>
|
||
|
||
<a id="1104" class="Comment">-- The following is then also an iso</a>
|
||
<a id="Distributive.distributeʳ"></a><a id="1143" href="Categories.Category.Distributive.html#1143" class="Function">distributeʳ</a> <a id="1155" class="Symbol">:</a> <a id="1157" class="Symbol">∀</a> <a id="1159" class="Symbol">{</a><a id="1160" href="Categories.Category.Distributive.html#1160" class="Bound">A</a> <a id="1162" href="Categories.Category.Distributive.html#1162" class="Bound">B</a> <a id="1164" href="Categories.Category.Distributive.html#1164" class="Bound">C</a> <a id="1166" class="Symbol">:</a> <a id="1168" href="Categories.Category.Core.html#559" class="Field">Obj</a><a id="1171" class="Symbol">}</a> <a id="1173" class="Symbol">→</a> <a id="1176" href="Categories.Category.Distributive.html#1162" class="Bound">B</a> <a id="1178" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="1180" href="Categories.Category.Distributive.html#1160" class="Bound">A</a> <a id="1182" href="Categories.Category.Cocartesian.html#1549" class="Function Operator">+</a> <a id="1184" href="Categories.Category.Distributive.html#1164" class="Bound">C</a> <a id="1186" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="1188" href="Categories.Category.Distributive.html#1160" class="Bound">A</a> <a id="1190" href="Categories.Category.Core.html#575" class="Field Operator">⇒</a> <a id="1192" class="Symbol">(</a><a id="1193" href="Categories.Category.Distributive.html#1162" class="Bound">B</a> <a id="1195" href="Categories.Category.Cocartesian.html#1549" class="Function Operator">+</a> <a id="1197" href="Categories.Category.Distributive.html#1164" class="Bound">C</a><a id="1198" class="Symbol">)</a> <a id="1200" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="1202" href="Categories.Category.Distributive.html#1160" class="Bound">A</a>
|
||
<a id="1206" href="Categories.Category.Distributive.html#1143" class="Function">distributeʳ</a> <a id="1218" class="Symbol">=</a> <a id="1220" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="1222" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a> <a id="1225" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1227" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1230" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="1232" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a> <a id="1235" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1237" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1240" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a>
|
||
|
||
<a id="Distributive.isIsoʳ"></a><a id="1245" href="Categories.Category.Distributive.html#1245" class="Function">isIsoʳ</a> <a id="1252" class="Symbol">:</a> <a id="1254" class="Symbol">∀</a> <a id="1256" class="Symbol">{</a><a id="1257" href="Categories.Category.Distributive.html#1257" class="Bound">A</a> <a id="1259" href="Categories.Category.Distributive.html#1259" class="Bound">B</a> <a id="1261" href="Categories.Category.Distributive.html#1261" class="Bound">C</a> <a id="1263" class="Symbol">:</a> <a id="1265" href="Categories.Category.Core.html#559" class="Field">Obj</a><a id="1268" class="Symbol">}</a> <a id="1270" class="Symbol">→</a> <a id="1273" href="Categories.Morphism.html#1826" class="Record">IsIso</a> <a id="1279" class="Symbol">(</a><a id="1280" href="Categories.Category.Distributive.html#1143" class="Function">distributeʳ</a> <a id="1292" class="Symbol">{</a><a id="1293" href="Categories.Category.Distributive.html#1257" class="Bound">A</a><a id="1294" class="Symbol">}</a> <a id="1296" class="Symbol">{</a><a id="1297" href="Categories.Category.Distributive.html#1259" class="Bound">B</a><a id="1298" class="Symbol">}</a> <a id="1300" class="Symbol">{</a><a id="1301" href="Categories.Category.Distributive.html#1261" class="Bound">C</a><a id="1302" class="Symbol">})</a>
|
||
<a id="1307" href="Categories.Category.Distributive.html#1245" class="Function">isIsoʳ</a> <a id="1314" class="Symbol">{</a><a id="1315" href="Categories.Category.Distributive.html#1315" class="Bound">A</a><a id="1316" class="Symbol">}</a> <a id="1318" class="Symbol">{</a><a id="1319" href="Categories.Category.Distributive.html#1319" class="Bound">B</a><a id="1320" class="Symbol">}</a> <a id="1322" class="Symbol">{</a><a id="1323" href="Categories.Category.Distributive.html#1323" class="Bound">C</a><a id="1324" class="Symbol">}</a> <a id="1326" class="Symbol">=</a> <a id="1328" class="Keyword">record</a>
|
||
<a id="1340" class="Symbol">{</a> <a id="1342" href="Categories.Morphism.html#1879" class="Field">inv</a> <a id="1346" class="Symbol">=</a> <a id="1348" class="Symbol">((</a><a id="1350" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1355" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="1358" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1362" class="Symbol">)</a> <a id="1364" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1366" href="Categories.Morphism.html#1879" class="Function">inv</a><a id="1369" class="Symbol">)</a> <a id="1371" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1373" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a>
|
||
<a id="1382" class="Symbol">;</a> <a id="1384" href="Categories.Morphism.html#1895" class="Field">iso</a> <a id="1388" class="Symbol">=</a> <a id="1390" class="Keyword">record</a>
|
||
<a id="1404" class="Symbol">{</a> <a id="1406" href="Categories.Morphism.html#1586" class="Field">isoˡ</a> <a id="1411" class="Symbol">=</a> <a id="1413" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="1428" class="Symbol">(((</a><a id="1431" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1436" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="1439" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1443" class="Symbol">)</a> <a id="1445" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1447" href="Categories.Morphism.html#1879" class="Function">inv</a><a id="1450" class="Symbol">)</a> <a id="1452" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1454" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1458" class="Symbol">)</a> <a id="1460" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1462" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="1464" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a> <a id="1467" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1469" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1472" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="1474" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a> <a id="1477" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1479" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1482" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a> <a id="1522" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="1525" href="Categories.Category.Cocartesian.html#2736" class="Function">∘[]</a> <a id="1529" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="1540" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="1542" class="Symbol">(((</a><a id="1545" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1550" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="1553" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1557" class="Symbol">)</a> <a id="1559" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1561" href="Categories.Morphism.html#1879" class="Function">inv</a><a id="1564" class="Symbol">)</a> <a id="1566" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1568" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1572" class="Symbol">)</a> <a id="1574" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1576" class="Symbol">(</a><a id="1577" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a> <a id="1580" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1582" href="Categories.Category.Core.html#630" class="Field">id</a><a id="1584" class="Symbol">)</a> <a id="1586" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="1588" class="Symbol">(((</a><a id="1591" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1596" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="1599" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1603" class="Symbol">)</a> <a id="1605" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1607" href="Categories.Morphism.html#1879" class="Function">inv</a><a id="1610" class="Symbol">)</a> <a id="1612" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1614" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1618" class="Symbol">)</a> <a id="1620" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1622" class="Symbol">(</a><a id="1623" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a> <a id="1626" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1628" href="Categories.Category.Core.html#630" class="Field">id</a><a id="1630" class="Symbol">)</a> <a id="1632" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a> <a id="1634" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="1637" href="Categories.Object.Coproduct.html#816" class="Function">[]-cong₂</a> <a id="1646" class="Symbol">(</a><a id="1647" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="1653" href="Categories.Category.BinaryProducts.html#4512" class="Function">swap∘⁂</a><a id="1659" class="Symbol">)</a> <a id="1661" class="Symbol">(</a><a id="1662" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="1668" href="Categories.Category.BinaryProducts.html#4512" class="Function">swap∘⁂</a><a id="1674" class="Symbol">)</a> <a id="1676" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="1686" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="1688" class="Symbol">((</a><a id="1690" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1695" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="1698" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1702" class="Symbol">)</a> <a id="1704" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1706" href="Categories.Morphism.html#1879" class="Function">inv</a><a id="1709" class="Symbol">)</a> <a id="1711" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1713" class="Symbol">(</a><a id="1714" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1717" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1719" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a><a id="1721" class="Symbol">)</a> <a id="1723" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1725" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1730" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="1732" class="Symbol">((</a><a id="1734" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1739" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="1742" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1746" class="Symbol">)</a> <a id="1748" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1750" href="Categories.Morphism.html#1879" class="Function">inv</a><a id="1753" class="Symbol">)</a> <a id="1755" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1757" class="Symbol">(</a><a id="1758" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1761" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1763" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a><a id="1765" class="Symbol">)</a> <a id="1767" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1769" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1774" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a> <a id="1780" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="1784" href="Categories.Category.Cocartesian.html#2736" class="Function">∘[]</a> <a id="1788" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="1798" class="Symbol">((</a><a id="1800" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1805" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="1808" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1812" class="Symbol">)</a> <a id="1814" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1816" href="Categories.Morphism.html#1879" class="Function">inv</a><a id="1819" class="Symbol">)</a> <a id="1821" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1823" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="1825" class="Symbol">(</a><a id="1826" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1829" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1831" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a><a id="1833" class="Symbol">)</a> <a id="1835" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1837" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1842" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="1844" class="Symbol">(</a><a id="1845" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1848" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1850" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a><a id="1852" class="Symbol">)</a> <a id="1854" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1856" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1861" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a> <a id="1892" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="1896" href="Categories.Category.Core.html#2734" class="Function Operator">refl⟩∘⟨</a> <a id="1904" href="Categories.Category.Cocartesian.html#2662" class="Function">[]∘+₁</a> <a id="1910" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="1920" class="Symbol">((</a><a id="1922" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1927" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="1930" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1934" class="Symbol">)</a> <a id="1936" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1938" href="Categories.Morphism.html#1879" class="Function">inv</a><a id="1941" class="Symbol">)</a> <a id="1943" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1945" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="1947" class="Symbol">(</a><a id="1948" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1951" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1953" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a><a id="1955" class="Symbol">)</a> <a id="1957" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="1959" class="Symbol">(</a><a id="1960" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="1963" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="1965" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a><a id="1967" class="Symbol">)</a> <a id="1969" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a> <a id="1971" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="1973" class="Symbol">(</a><a id="1974" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="1979" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="1982" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="1986" class="Symbol">)</a> <a id="2014" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2017" href="Categories.Morphism.Reasoning.Core.html#7037" class="Function">cancelInner</a> <a id="2029" href="Categories.Morphism.html#1586" class="Function">isoˡ</a> <a id="2034" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2044" class="Symbol">(</a><a id="2045" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2050" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="2053" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="2057" class="Symbol">)</a> <a id="2059" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2061" class="Symbol">(</a><a id="2062" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2067" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="2070" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="2074" class="Symbol">)</a> <a id="2138" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2141" href="Categories.Category.Cocartesian.html#2699" class="Function">+₁∘+₁</a> <a id="2147" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2157" class="Symbol">(</a><a id="2158" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2163" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2165" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="2169" class="Symbol">)</a> <a id="2171" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="2174" class="Symbol">(</a><a id="2175" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2180" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2182" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="2186" class="Symbol">)</a> <a id="2251" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2254" href="Categories.Category.Cocartesian.html#2622" class="Function">+₁-cong₂</a> <a id="2263" href="Categories.Category.BinaryProducts.html#4688" class="Function">swap∘swap</a> <a id="2273" href="Categories.Category.BinaryProducts.html#4688" class="Function">swap∘swap</a> <a id="2283" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2293" class="Symbol">(</a><a id="2294" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="2297" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="2300" href="Categories.Category.Core.html#630" class="Field">id</a><a id="2302" class="Symbol">)</a> <a id="2387" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2390" href="Categories.Category.Cocartesian.html#1714" class="Function">+-unique</a> <a id="2399" href="Categories.Morphism.Reasoning.Core.html#1309" class="Function">id-comm-sym</a> <a id="2411" href="Categories.Morphism.Reasoning.Core.html#1309" class="Function">id-comm-sym</a> <a id="2423" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2433" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="2527" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
<a id="2536" class="Symbol">;</a> <a id="2538" href="Categories.Morphism.html#1612" class="Field">isoʳ</a> <a id="2543" class="Symbol">=</a> <a id="2545" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="2560" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="2562" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a> <a id="2565" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="2567" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="2570" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="2572" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a> <a id="2575" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="2577" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="2580" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a> <a id="2582" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2584" class="Symbol">((</a><a id="2586" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2591" href="Categories.Category.Cocartesian.html#2133" class="Function Operator">+₁</a> <a id="2594" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a><a id="2598" class="Symbol">)</a> <a id="2600" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2602" href="Categories.Morphism.html#1879" class="Function">inv</a><a id="2605" class="Symbol">)</a> <a id="2607" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2609" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2615" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2618" href="Categories.Morphism.Reasoning.Core.html#7730" class="Function">pull-first</a> <a id="2629" href="Categories.Category.Cocartesian.html#2662" class="Function">[]∘+₁</a> <a id="2635" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2645" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="2647" class="Symbol">(</a><a id="2648" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a> <a id="2651" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="2653" href="Categories.Category.Core.html#630" class="Field">id</a><a id="2655" class="Symbol">)</a> <a id="2657" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2659" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2664" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="2666" class="Symbol">(</a><a id="2667" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a> <a id="2670" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="2672" href="Categories.Category.Core.html#630" class="Field">id</a><a id="2674" class="Symbol">)</a> <a id="2676" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2678" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2683" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a> <a id="2685" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2687" href="Categories.Morphism.html#1879" class="Function">inv</a> <a id="2691" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2693" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2700" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="2704" href="Categories.Object.Coproduct.html#816" class="Function">[]-cong₂</a> <a id="2713" href="Categories.Category.BinaryProducts.html#4512" class="Function">swap∘⁂</a> <a id="2720" href="Categories.Category.BinaryProducts.html#4512" class="Function">swap∘⁂</a> <a id="2727" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="2735" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="2745" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="2747" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2752" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2754" class="Symbol">(</a><a id="2755" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="2758" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="2760" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a><a id="2762" class="Symbol">)</a> <a id="2764" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="2766" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2771" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2773" class="Symbol">(</a><a id="2774" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="2777" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="2779" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a><a id="2781" class="Symbol">)</a> <a id="2783" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a> <a id="2785" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2787" href="Categories.Morphism.html#1879" class="Function">inv</a> <a id="2791" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2793" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2800" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="2804" href="Categories.Category.Cocartesian.html#2736" class="Function">∘[]</a> <a id="2808" href="Categories.Category.Core.html#2837" class="Function Operator">⟩∘⟨refl</a> <a id="2816" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">⟩</a>
|
||
<a id="2826" class="Symbol">(</a><a id="2827" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2832" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2834" href="Categories.Object.Coproduct.html#532" class="Function Operator">[</a> <a id="2836" class="Symbol">(</a><a id="2837" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="2840" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="2842" href="Categories.Object.Coproduct.html#492" class="Function">i₁</a><a id="2844" class="Symbol">)</a> <a id="2846" href="Categories.Object.Coproduct.html#532" class="Function Operator">,</a> <a id="2848" class="Symbol">(</a><a id="2849" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="2852" href="Categories.Category.BinaryProducts.html#1465" class="Function Operator">⁂</a> <a id="2854" href="Categories.Object.Coproduct.html#512" class="Function">i₂</a><a id="2856" class="Symbol">)</a> <a id="2858" href="Categories.Object.Coproduct.html#532" class="Function Operator">]</a><a id="2859" class="Symbol">)</a> <a id="2861" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2863" href="Categories.Morphism.html#1879" class="Function">inv</a> <a id="2867" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2869" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2881" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2884" href="Categories.Morphism.Reasoning.Core.html#7037" class="Function">cancelInner</a> <a id="2896" href="Categories.Morphism.html#1612" class="Function">isoʳ</a> <a id="2902" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2912" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2917" href="Categories.Category.Core.html#656" class="Field Operator">∘</a> <a id="2919" href="Categories.Category.BinaryProducts.html#2048" class="Function">swap</a> <a id="2967" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2970" href="Categories.Category.BinaryProducts.html#4688" class="Function">swap∘swap</a> <a id="2980" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2990" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="3045" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
<a id="3054" class="Symbol">}</a>
|
||
<a id="3061" class="Symbol">}</a>
|
||
<a id="3067" class="Keyword">where</a>
|
||
<a id="3079" class="Keyword">open</a> <a id="3084" href="Categories.Morphism.html#1826" class="Module">IsIso</a> <a id="3090" class="Symbol">(</a><a id="3091" href="Categories.Category.Distributive.html#1042" class="Field">isIsoˡ</a> <a id="3098" class="Symbol">{</a><a id="3099" href="Categories.Category.Distributive.html#1315" class="Bound">A</a><a id="3100" class="Symbol">}</a> <a id="3102" class="Symbol">{</a><a id="3103" href="Categories.Category.Distributive.html#1319" class="Bound">B</a><a id="3104" class="Symbol">}</a> <a id="3106" class="Symbol">{</a><a id="3107" href="Categories.Category.Distributive.html#1323" class="Bound">C</a><a id="3108" class="Symbol">})</a>
|
||
|
||
<a id="3114" class="Comment">-- The inverse is what one is usually interested in</a>
|
||
<a id="Distributive.distributeˡ⁻¹"></a><a id="3168" href="Categories.Category.Distributive.html#3168" class="Function">distributeˡ⁻¹</a> <a id="3182" class="Symbol">:</a> <a id="3184" class="Symbol">∀</a> <a id="3186" class="Symbol">{</a><a id="3187" href="Categories.Category.Distributive.html#3187" class="Bound">A</a> <a id="3189" href="Categories.Category.Distributive.html#3189" class="Bound">B</a> <a id="3191" href="Categories.Category.Distributive.html#3191" class="Bound">C</a> <a id="3193" class="Symbol">:</a> <a id="3195" href="Categories.Category.Core.html#559" class="Field">Obj</a><a id="3198" class="Symbol">}</a> <a id="3200" class="Symbol">→</a> <a id="3202" href="Categories.Category.Distributive.html#3187" class="Bound">A</a> <a id="3204" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="3206" class="Symbol">(</a><a id="3207" href="Categories.Category.Distributive.html#3189" class="Bound">B</a> <a id="3209" href="Categories.Category.Cocartesian.html#1549" class="Function Operator">+</a> <a id="3211" href="Categories.Category.Distributive.html#3191" class="Bound">C</a><a id="3212" class="Symbol">)</a> <a id="3214" href="Categories.Category.Core.html#575" class="Field Operator">⇒</a> <a id="3216" href="Categories.Category.Distributive.html#3187" class="Bound">A</a> <a id="3218" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="3220" href="Categories.Category.Distributive.html#3189" class="Bound">B</a> <a id="3222" href="Categories.Category.Cocartesian.html#1549" class="Function Operator">+</a> <a id="3224" href="Categories.Category.Distributive.html#3187" class="Bound">A</a> <a id="3226" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="3228" href="Categories.Category.Distributive.html#3191" class="Bound">C</a>
|
||
<a id="3232" href="Categories.Category.Distributive.html#3168" class="Function">distributeˡ⁻¹</a> <a id="3246" class="Symbol">=</a> <a id="3248" href="Categories.Morphism.html#1879" class="Field">IsIso.inv</a> <a id="3258" href="Categories.Category.Distributive.html#1042" class="Field">isIsoˡ</a>
|
||
|
||
<a id="Distributive.distributeʳ⁻¹"></a><a id="3272" href="Categories.Category.Distributive.html#3272" class="Function">distributeʳ⁻¹</a> <a id="3286" class="Symbol">:</a> <a id="3288" class="Symbol">∀</a> <a id="3290" class="Symbol">{</a><a id="3291" href="Categories.Category.Distributive.html#3291" class="Bound">A</a> <a id="3293" href="Categories.Category.Distributive.html#3293" class="Bound">B</a> <a id="3295" href="Categories.Category.Distributive.html#3295" class="Bound">C</a> <a id="3297" class="Symbol">:</a> <a id="3299" href="Categories.Category.Core.html#559" class="Field">Obj</a><a id="3302" class="Symbol">}</a> <a id="3304" class="Symbol">→</a> <a id="3306" class="Symbol">(</a><a id="3307" href="Categories.Category.Distributive.html#3293" class="Bound">B</a> <a id="3309" href="Categories.Category.Cocartesian.html#1549" class="Function Operator">+</a> <a id="3311" href="Categories.Category.Distributive.html#3295" class="Bound">C</a><a id="3312" class="Symbol">)</a> <a id="3314" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="3316" href="Categories.Category.Distributive.html#3291" class="Bound">A</a> <a id="3318" href="Categories.Category.Core.html#575" class="Field Operator">⇒</a> <a id="3320" href="Categories.Category.Distributive.html#3293" class="Bound">B</a> <a id="3322" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="3324" href="Categories.Category.Distributive.html#3291" class="Bound">A</a> <a id="3326" href="Categories.Category.Cocartesian.html#1549" class="Function Operator">+</a> <a id="3328" href="Categories.Category.Distributive.html#3295" class="Bound">C</a> <a id="3330" href="Categories.Category.BinaryProducts.html#1053" class="Function Operator">×</a> <a id="3332" href="Categories.Category.Distributive.html#3291" class="Bound">A</a>
|
||
<a id="3336" href="Categories.Category.Distributive.html#3272" class="Function">distributeʳ⁻¹</a> <a id="3350" class="Symbol">=</a> <a id="3352" href="Categories.Morphism.html#1879" class="Field">IsIso.inv</a> <a id="3362" href="Categories.Category.Distributive.html#1245" class="Function">isIsoʳ</a>
|
||
</pre></body></html> |