bsc-leon-vatthauer/agda/bsc-thesis/Categories.Category.Construction.F-Algebras.html

116 lines
39 KiB
HTML
Raw Normal View History

2024-02-09 17:53:52 +01:00
<!DOCTYPE HTML>
<html><head><meta charset="utf-8"><title>Categories.Category.Construction.F-Algebras</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">module</a> <a id="43" href="Categories.Category.Construction.F-Algebras.html" class="Module">Categories.Category.Construction.F-Algebras</a> <a id="87" class="Keyword">where</a>
<a id="94" class="Keyword">open</a> <a id="99" class="Keyword">import</a> <a id="106" href="Level.html" class="Module">Level</a>
<a id="112" class="Keyword">open</a> <a id="117" class="Keyword">import</a> <a id="124" href="Data.Product.html" class="Module">Data.Product</a> <a id="137" class="Keyword">using</a> <a id="143" class="Symbol">(</a><a id="144" href="Data.Product.Base.html#636" class="Field">proj₁</a><a id="149" class="Symbol">;</a> <a id="151" href="Data.Product.Base.html#650" class="Field">proj₂</a><a id="156" class="Symbol">)</a>
<a id="159" class="Keyword">open</a> <a id="164" class="Keyword">import</a> <a id="171" href="Categories.Category.html" class="Module">Categories.Category</a>
<a id="191" class="Keyword">open</a> <a id="196" class="Keyword">import</a> <a id="203" href="Categories.Functor.html" class="Module">Categories.Functor</a> <a id="222" class="Keyword">hiding</a> <a id="229" class="Symbol">(</a><a id="230" href="Categories.Functor.html#349" class="Function">id</a><a id="232" class="Symbol">)</a>
<a id="234" class="Keyword">open</a> <a id="239" class="Keyword">import</a> <a id="246" href="Categories.Functor.Algebra.html" class="Module">Categories.Functor.Algebra</a>
<a id="273" class="Keyword">open</a> <a id="278" class="Keyword">import</a> <a id="285" href="Categories.Object.Initial.html" class="Module">Categories.Object.Initial</a>
<a id="311" class="Keyword">import</a> <a id="318" href="Categories.Morphism.Reasoning.html" class="Module">Categories.Morphism.Reasoning</a> <a id="348" class="Symbol">as</a> <a id="351" class="Module">MR</a>
<a id="354" class="Keyword">import</a> <a id="361" href="Categories.Morphism.html" class="Module">Categories.Morphism</a> <a id="381" class="Symbol">as</a> <a id="384" class="Module">Mor</a> <a id="388" class="Keyword">using</a> <a id="394" class="Symbol">(</a><a id="395" href="Categories.Morphism.html#1958" class="Record Operator">_≅_</a><a id="398" class="Symbol">)</a>
<a id="401" class="Keyword">private</a>
<a id="411" class="Keyword">variable</a>
<a id="424" href="Categories.Category.Construction.F-Algebras.html#424" class="Generalizable">o</a> <a id="426" href="Categories.Category.Construction.F-Algebras.html#426" class="Generalizable"></a> <a id="428" href="Categories.Category.Construction.F-Algebras.html#428" class="Generalizable">e</a> <a id="430" class="Symbol">:</a> <a id="432" href="Agda.Primitive.html#742" class="Postulate">Level</a>
<a id="442" href="Categories.Category.Construction.F-Algebras.html#442" class="Generalizable">𝒞</a> <a id="444" class="Symbol">:</a> <a id="446" href="Categories.Category.Core.html#442" class="Record">Category</a> <a id="455" href="Categories.Category.Construction.F-Algebras.html#424" class="Generalizable">o</a> <a id="457" href="Categories.Category.Construction.F-Algebras.html#426" class="Generalizable"></a> <a id="459" href="Categories.Category.Construction.F-Algebras.html#428" class="Generalizable">e</a>
<a id="F-Algebras"></a><a id="462" href="Categories.Category.Construction.F-Algebras.html#462" class="Function">F-Algebras</a> <a id="473" class="Symbol">:</a> <a id="475" class="Symbol">{</a><a id="476" href="Categories.Category.Construction.F-Algebras.html#476" class="Bound">𝒞</a> <a id="478" class="Symbol">:</a> <a id="480" href="Categories.Category.Core.html#442" class="Record">Category</a> <a id="489" href="Categories.Category.Construction.F-Algebras.html#424" class="Generalizable">o</a> <a id="491" href="Categories.Category.Construction.F-Algebras.html#426" class="Generalizable"></a> <a id="493" href="Categories.Category.Construction.F-Algebras.html#428" class="Generalizable">e</a><a id="494" class="Symbol">}</a> <a id="496" class="Symbol"></a> <a id="498" href="Categories.Functor.html#283" class="Function">Endofunctor</a> <a id="510" href="Categories.Category.Construction.F-Algebras.html#476" class="Bound">𝒞</a> <a id="512" class="Symbol"></a> <a id="514" href="Categories.Category.Core.html#442" class="Record">Category</a> <a id="523" class="Symbol">(</a><a id="524" href="Categories.Category.Construction.F-Algebras.html#426" class="Generalizable"></a> <a id="526" href="Agda.Primitive.html#961" class="Primitive Operator"></a> <a id="528" href="Categories.Category.Construction.F-Algebras.html#424" class="Generalizable">o</a><a id="529" class="Symbol">)</a> <a id="531" class="Symbol">(</a><a id="532" href="Categories.Category.Construction.F-Algebras.html#428" class="Generalizable">e</a> <a id="534" href="Agda.Primitive.html#961" class="Primitive Operator"></a> <a id="536" href="Categories.Category.Construction.F-Algebras.html#426" class="Generalizable"></a><a id="537" class="Symbol">)</a> <a id="539" href="Categories.Category.Construction.F-Algebras.html#428" class="Generalizable">e</a>
<a id="541" href="Categories.Category.Construction.F-Algebras.html#462" class="Function">F-Algebras</a> <a id="552" class="Symbol">{</a><a id="553" class="Argument">𝒞</a> <a id="555" class="Symbol">=</a> <a id="557" href="Categories.Category.Construction.F-Algebras.html#557" class="Bound">𝒞</a><a id="558" class="Symbol">}</a> <a id="560" href="Categories.Category.Construction.F-Algebras.html#560" class="Bound">F</a> <a id="562" class="Symbol">=</a> <a id="564" class="Keyword">record</a>
<a id="573" class="Symbol">{</a> <a id="575" href="Categories.Category.Core.html#559" class="Field">Obj</a> <a id="585" class="Symbol">=</a> <a id="587" href="Categories.Functor.Algebra.html#487" class="Record">F-Algebra</a> <a id="597" href="Categories.Category.Construction.F-Algebras.html#560" class="Bound">F</a>
<a id="601" class="Symbol">;</a> <a id="603" href="Categories.Category.Core.html#575" class="Field Operator">_⇒_</a> <a id="613" class="Symbol">=</a> <a id="615" href="Categories.Functor.Algebra.html#1307" class="Record">F-Algebra-Morphism</a>
<a id="636" class="Symbol">;</a> <a id="638" href="Categories.Category.Core.html#595" class="Field Operator">_≈_</a> <a id="648" class="Symbol">=</a> <a id="650" class="Symbol">λ</a> <a id="652" href="Categories.Category.Construction.F-Algebras.html#652" class="Bound">α₁</a> <a id="655" href="Categories.Category.Construction.F-Algebras.html#655" class="Bound">α₂</a> <a id="658" class="Symbol"></a> <a id="660" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="662" href="Categories.Category.Construction.F-Algebras.html#652" class="Bound">α₁</a> <a id="665" href="Categories.Category.Core.html#595" class="Function Operator"></a> <a id="667" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="669" href="Categories.Category.Construction.F-Algebras.html#655" class="Bound">α₂</a>
<a id="674" class="Symbol">;</a> <a id="676" href="Categories.Category.Core.html#656" class="Field Operator">_∘_</a> <a id="686" class="Symbol">=</a> <a id="688" class="Symbol">λ</a> <a id="690" href="Categories.Category.Construction.F-Algebras.html#690" class="Bound">α₁</a> <a id="693" href="Categories.Category.Construction.F-Algebras.html#693" class="Bound">α₂</a> <a id="696" class="Symbol"></a> <a id="698" class="Keyword">record</a> <a id="705" class="Symbol">{</a> <a id="707" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="709" class="Symbol">=</a> <a id="711" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="713" href="Categories.Category.Construction.F-Algebras.html#690" class="Bound">α₁</a> <a id="716" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="718" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="720" href="Categories.Category.Construction.F-Algebras.html#693" class="Bound">α₂</a> <a id="723" class="Symbol">;</a> <a id="725" href="Categories.Functor.Algebra.html#1388" class="Field">commutes</a> <a id="734" class="Symbol">=</a> <a id="736" href="Categories.Category.Construction.F-Algebras.html#1225" class="Function">commut</a> <a id="743" href="Categories.Category.Construction.F-Algebras.html#690" class="Bound">α₁</a> <a id="746" href="Categories.Category.Construction.F-Algebras.html#693" class="Bound">α₂</a> <a id="749" class="Symbol">}</a>
<a id="753" class="Symbol">;</a> <a id="755" href="Categories.Category.Core.html#630" class="Field">id</a> <a id="765" class="Symbol">=</a> <a id="767" class="Keyword">record</a> <a id="774" class="Symbol">{</a> <a id="776" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="778" class="Symbol">=</a> <a id="780" href="Categories.Category.Core.html#630" class="Function">id</a> <a id="783" class="Symbol">;</a> <a id="785" href="Categories.Functor.Algebra.html#1388" class="Field">commutes</a> <a id="794" class="Symbol">=</a> <a id="796" href="Categories.Morphism.Reasoning.Core.html#1309" class="Function">id-comm-sym</a> <a id="808" href="Categories.Category.Core.html#3061" class="Function Operator"></a> <a id="810" href="Categories.Category.Core.html#3005" class="Function"></a> <a id="812" class="Symbol">(</a><a id="813" href="Categories.Category.Core.html#1706" class="Function">∘-resp-≈ʳ</a> <a id="823" href="Categories.Functor.Core.html#511" class="Field">identity</a><a id="831" class="Symbol">)</a> <a id="833" class="Symbol">}</a>
<a id="837" class="Symbol">;</a> <a id="839" href="Categories.Category.Core.html#715" class="Field">assoc</a> <a id="849" class="Symbol">=</a> <a id="851" href="Categories.Category.Core.html#715" class="Function">assoc</a>
<a id="859" class="Symbol">;</a> <a id="861" href="Categories.Category.Core.html#1004" class="Field">sym-assoc</a> <a id="871" class="Symbol">=</a> <a id="873" href="Categories.Category.Core.html#1004" class="Function">sym-assoc</a>
<a id="885" class="Symbol">;</a> <a id="887" href="Categories.Category.Core.html#1096" class="Field">identityˡ</a> <a id="897" class="Symbol">=</a> <a id="899" href="Categories.Category.Core.html#1096" class="Function">identityˡ</a>
<a id="911" class="Symbol">;</a> <a id="913" href="Categories.Category.Core.html#1145" class="Field">identityʳ</a> <a id="923" class="Symbol">=</a> <a id="925" href="Categories.Category.Core.html#1145" class="Function">identityʳ</a>
<a id="937" class="Symbol">;</a> <a id="939" href="Categories.Category.Core.html#1339" class="Field">identity²</a> <a id="949" class="Symbol">=</a> <a id="951" href="Categories.Category.Core.html#1339" class="Function">identity²</a>
<a id="963" class="Symbol">;</a> <a id="965" href="Categories.Category.Core.html#1384" class="Field">equiv</a> <a id="975" class="Symbol">=</a> <a id="977" class="Keyword">record</a>
<a id="988" class="Symbol">{</a> <a id="990" href="Relation.Binary.Structures.html#1596" class="Field">refl</a> <a id="996" class="Symbol">=</a> <a id="998" href="Relation.Binary.Structures.html#1596" class="Function">refl</a>
<a id="1007" class="Symbol">;</a> <a id="1009" href="Relation.Binary.Structures.html#1622" class="Field">sym</a> <a id="1015" class="Symbol">=</a> <a id="1017" href="Relation.Binary.Structures.html#1622" class="Function">sym</a>
<a id="1025" class="Symbol">;</a> <a id="1027" href="Relation.Binary.Structures.html#1648" class="Field">trans</a> <a id="1033" class="Symbol">=</a> <a id="1035" href="Relation.Binary.Structures.html#1648" class="Function">trans</a>
<a id="1045" class="Symbol">}</a>
<a id="1049" class="Symbol">;</a> <a id="1051" href="Categories.Category.Core.html#1438" class="Field">∘-resp-≈</a> <a id="1061" class="Symbol">=</a> <a id="1063" href="Categories.Category.Core.html#1438" class="Function">∘-resp-≈</a>
<a id="1074" class="Symbol">}</a>
<a id="1078" class="Keyword">where</a>
<a id="1088" class="Keyword">open</a> <a id="1093" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="1102" href="Categories.Category.Construction.F-Algebras.html#557" class="Bound">𝒞</a>
<a id="1108" class="Keyword">open</a> <a id="1113" href="Categories.Morphism.Reasoning.html" class="Module">MR</a> <a id="1116" href="Categories.Category.Construction.F-Algebras.html#557" class="Bound">𝒞</a>
<a id="1122" class="Keyword">open</a> <a id="1127" href="Categories.Category.Core.html#2462" class="Module">HomReasoning</a>
<a id="1144" class="Keyword">open</a> <a id="1149" href="Categories.Category.Core.html#1530" class="Module">Equiv</a>
<a id="1159" class="Keyword">open</a> <a id="1164" href="Categories.Functor.Core.html#248" class="Module">Functor</a> <a id="1172" href="Categories.Category.Construction.F-Algebras.html#560" class="Bound">F</a>
<a id="1178" class="Keyword">open</a> <a id="1183" href="Categories.Functor.Algebra.html#1307" class="Module">F-Algebra-Morphism</a>
<a id="1206" class="Keyword">open</a> <a id="1211" href="Categories.Functor.Algebra.html#487" class="Module">F-Algebra</a>
<a id="1225" href="Categories.Category.Construction.F-Algebras.html#1225" class="Function">commut</a> <a id="1232" class="Symbol">:</a> <a id="1234" class="Symbol">{</a><a id="1235" href="Categories.Category.Construction.F-Algebras.html#1235" class="Bound">A</a> <a id="1237" href="Categories.Category.Construction.F-Algebras.html#1237" class="Bound">B</a> <a id="1239" href="Categories.Category.Construction.F-Algebras.html#1239" class="Bound">C</a> <a id="1241" class="Symbol">:</a> <a id="1243" href="Categories.Functor.Algebra.html#487" class="Record">F-Algebra</a> <a id="1253" href="Categories.Category.Construction.F-Algebras.html#560" class="Bound">F</a><a id="1254" class="Symbol">}</a> <a id="1256" class="Symbol">(</a><a id="1257" href="Categories.Category.Construction.F-Algebras.html#1257" class="Bound">α₁</a> <a id="1260" class="Symbol">:</a> <a id="1262" href="Categories.Functor.Algebra.html#1307" class="Record">F-Algebra-Morphism</a> <a id="1281" href="Categories.Category.Construction.F-Algebras.html#1237" class="Bound">B</a> <a id="1283" href="Categories.Category.Construction.F-Algebras.html#1239" class="Bound">C</a><a id="1284" class="Symbol">)</a> <a id="1286" class="Symbol">(</a><a id="1287" href="Categories.Category.Construction.F-Algebras.html#1287" class="Bound">α₂</a> <a id="1290" class="Symbol">:</a> <a id="1292" href="Categories.Functor.Algebra.html#1307" class="Record">F-Algebra-Morphism</a> <a id="1311" href="Categories.Category.Construction.F-Algebras.html#1235" class="Bound">A</a> <a id="1313" href="Categories.Category.Construction.F-Algebras.html#1237" class="Bound">B</a><a id="1314" class="Symbol">)</a> <a id="1316" class="Symbol"></a>
<a id="1324" class="Symbol">(</a><a id="1325" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1327" href="Categories.Category.Construction.F-Algebras.html#1257" class="Bound">α₁</a> <a id="1330" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1332" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1334" href="Categories.Category.Construction.F-Algebras.html#1287" class="Bound">α₂</a><a id="1336" class="Symbol">)</a> <a id="1338" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1340" href="Categories.Functor.Algebra.html#587" class="Field">α</a> <a id="1342" href="Categories.Category.Construction.F-Algebras.html#1235" class="Bound">A</a> <a id="1344" href="Categories.Category.Core.html#595" class="Function Operator"></a> <a id="1346" href="Categories.Functor.Algebra.html#587" class="Field">α</a> <a id="1348" href="Categories.Category.Construction.F-Algebras.html#1239" class="Bound">C</a> <a id="1350" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1352" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1355" class="Symbol">(</a><a id="1356" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1358" href="Categories.Category.Construction.F-Algebras.html#1257" class="Bound">α₁</a> <a id="1361" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1363" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1365" href="Categories.Category.Construction.F-Algebras.html#1287" class="Bound">α₂</a><a id="1367" class="Symbol">)</a>
<a id="1373" href="Categories.Category.Construction.F-Algebras.html#1225" class="Function">commut</a> <a id="1380" class="Symbol">{</a><a id="1381" href="Categories.Category.Construction.F-Algebras.html#1381" class="Bound">A</a><a id="1382" class="Symbol">}</a> <a id="1384" class="Symbol">{</a><a id="1385" href="Categories.Category.Construction.F-Algebras.html#1385" class="Bound">B</a><a id="1386" class="Symbol">}</a> <a id="1388" class="Symbol">{</a><a id="1389" href="Categories.Category.Construction.F-Algebras.html#1389" class="Bound">C</a><a id="1390" class="Symbol">}</a> <a id="1392" href="Categories.Category.Construction.F-Algebras.html#1392" class="Bound">α₁</a> <a id="1395" href="Categories.Category.Construction.F-Algebras.html#1395" class="Bound">α₂</a> <a id="1398" class="Symbol">=</a> <a id="1400" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="1412" class="Symbol">(</a><a id="1413" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1415" href="Categories.Category.Construction.F-Algebras.html#1392" class="Bound">α₁</a> <a id="1418" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1420" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1422" href="Categories.Category.Construction.F-Algebras.html#1395" class="Bound">α₂</a><a id="1424" class="Symbol">)</a> <a id="1426" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1428" href="Categories.Functor.Algebra.html#587" class="Field">α</a> <a id="1430" href="Categories.Category.Construction.F-Algebras.html#1381" class="Bound">A</a> <a id="1443" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="1446" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="1452" class="Symbol">(</a><a id="1453" href="Categories.Functor.Algebra.html#1388" class="Field">commutes</a> <a id="1462" href="Categories.Category.Construction.F-Algebras.html#1395" class="Bound">α₂</a><a id="1464" class="Symbol">)</a> <a id="1466" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="1474" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1476" href="Categories.Category.Construction.F-Algebras.html#1392" class="Bound">α₁</a> <a id="1479" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1481" class="Symbol">(</a><a id="1482" href="Categories.Functor.Algebra.html#587" class="Field">α</a> <a id="1484" href="Categories.Category.Construction.F-Algebras.html#1385" class="Bound">B</a> <a id="1486" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1488" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1491" class="Symbol">(</a><a id="1492" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1494" href="Categories.Category.Construction.F-Algebras.html#1395" class="Bound">α₂</a><a id="1496" class="Symbol">))</a> <a id="1505" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="1508" href="Categories.Morphism.Reasoning.Core.html#2048" class="Function">pullˡ</a> <a id="1514" class="Symbol">(</a><a id="1515" href="Categories.Functor.Algebra.html#1388" class="Field">commutes</a> <a id="1524" href="Categories.Category.Construction.F-Algebras.html#1392" class="Bound">α₁</a><a id="1526" class="Symbol">)</a> <a id="1528" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="1536" class="Symbol">(</a><a id="1537" href="Categories.Functor.Algebra.html#587" class="Field">α</a> <a id="1539" href="Categories.Category.Construction.F-Algebras.html#1389" class="Bound">C</a> <a id="1541" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1543" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1546" class="Symbol">(</a><a id="1547" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1549" href="Categories.Category.Construction.F-Algebras.html#1392" class="Bound">α₁</a><a id="1551" class="Symbol">))</a> <a id="1554" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1556" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1559" class="Symbol">(</a><a id="1560" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1562" href="Categories.Category.Construction.F-Algebras.html#1395" class="Bound">α₂</a><a id="1564" class="Symbol">)</a> <a id="1567" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="1570" href="Categories.Morphism.Reasoning.Core.html#1914" class="Function">pullʳ</a> <a id="1576" class="Symbol">(</a><a id="1577" href="Categories.Category.Core.html#3005" class="Function"></a> <a id="1579" href="Categories.Functor.Core.html#565" class="Field">homomorphism</a><a id="1591" class="Symbol">)</a> <a id="1593" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="1601" href="Categories.Functor.Algebra.html#587" class="Field">α</a> <a id="1603" href="Categories.Category.Construction.F-Algebras.html#1389" class="Bound">C</a> <a id="1605" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1607" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1610" class="Symbol">(</a><a id="1611" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1613" href="Categories.Category.Construction.F-Algebras.html#1392" class="Bound">α₁</a> <a id="1616" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="1618" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="1620" href="Categories.Category.Construction.F-Algebras.html#1395" class="Bound">α₂</a><a id="1622" class="Symbol">)</a> <a id="1626" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="1630" class="Keyword">module</a> <a id="Lambek"></a><a id="1637" href="Categories.Category.Construction.F-Algebras.html#1637" class="Module">Lambek</a> <a id="1644" class="Symbol">{</a><a id="1645" href="Categories.Category.Construction.F-Algebras.html#1645" class="Bound">𝒞</a> <a id="1647" class="Symbol">:</a> <a id="1649" href="Categories.Category.Core.html#442" class="Record">Category</a> <a id="1658" href="Categories.Category.Construction.F-Algebras.html#424" class="Generalizable">o</a> <a id="1660" href="Categories.Category.Construction.F-Algebras.html#426" class="Generalizable"></a> <a id="1662" href="Categories.Category.Construction.F-Algebras.html#428" class="Generalizable">e</a><a id="1663" class="Symbol">}</a> <a id="1665" class="Symbol">{</a><a id="1666" href="Categories.Category.Construction.F-Algebras.html#1666" class="Bound">F</a> <a id="1668" class="Symbol">:</a> <a id="1670" href="Categories.Functor.html#283" class="Function">Endofunctor</a> <a id="1682" href="Categories.Category.Construction.F-Algebras.html#1645" class="Bound">𝒞</a><a id="1683" class="Symbol">}</a> <a id="1685" class="Symbol">(</a><a id="1686" href="Categories.Category.Construction.F-Algebras.html#1686" class="Bound">I</a> <a id="1688" class="Symbol">:</a> <a id="1690" href="Categories.Object.Initial.html#760" class="Record">Initial</a> <a id="1698" class="Symbol">(</a><a id="1699" href="Categories.Category.Construction.F-Algebras.html#462" class="Function">F-Algebras</a> <a id="1710" href="Categories.Category.Construction.F-Algebras.html#1666" class="Bound">F</a><a id="1711" class="Symbol">))</a> <a id="1714" class="Keyword">where</a>
<a id="1722" class="Keyword">open</a> <a id="1727" href="Categories.Category.Core.html#442" class="Module">Category</a> <a id="1736" href="Categories.Category.Construction.F-Algebras.html#1645" class="Bound">𝒞</a>
<a id="1740" class="Keyword">open</a> <a id="1745" href="Categories.Category.html#824" class="Module">Definitions</a> <a id="1757" href="Categories.Category.Construction.F-Algebras.html#1645" class="Bound">𝒞</a>
<a id="1761" class="Keyword">open</a> <a id="1766" href="Categories.Functor.Core.html#248" class="Module">Functor</a> <a id="1774" href="Categories.Category.Construction.F-Algebras.html#1666" class="Bound">F</a>
<a id="1778" class="Keyword">open</a> <a id="1783" href="Categories.Functor.Algebra.html#487" class="Module">F-Algebra</a> <a id="1793" class="Keyword">using</a> <a id="1799" class="Symbol">(</a><a id="1800" href="Categories.Functor.Algebra.html#587" class="Field">α</a><a id="1801" class="Symbol">)</a>
<a id="1806" class="Keyword">open</a> <a id="1811" href="Categories.Morphism.Reasoning.html" class="Module">MR</a> <a id="1814" href="Categories.Category.Construction.F-Algebras.html#1645" class="Bound">𝒞</a> <a id="1816" class="Keyword">using</a> <a id="1822" class="Symbol">(</a><a id="1823" href="Categories.Morphism.Reasoning.Core.html#4290" class="Function">glue</a><a id="1827" class="Symbol">)</a>
<a id="1831" class="Keyword">open</a> <a id="1836" href="Categories.Morphism.html" class="Module">Mor</a> <a id="1840" href="Categories.Category.Construction.F-Algebras.html#1645" class="Bound">𝒞</a>
<a id="1844" class="Keyword">open</a> <a id="1849" href="Categories.Object.Initial.html#760" class="Module">Initial</a> <a id="1857" href="Categories.Category.Construction.F-Algebras.html#1686" class="Bound">I</a> <a id="1859" class="Comment">-- so ⊥ is an F-Algebra, which is initial</a>
<a id="1904" class="Comment">-- While an expert might be able to decipher the proof at the nLab</a>
<a id="1973" class="Comment">-- (https://ncatlab.org/nlab/show/initial+algebra+of+an+endofunctor)</a>
<a id="2046" class="Comment">-- I (JC) have found that the notes at</a>
<a id="2087" class="Comment">-- http://www.cs.ru.nl/~jrot/coalgebra/ak-algebras.pdf</a>
<a id="2146" class="Comment">-- are easier to follow, and lead to the full proof below.</a>
<a id="2207" class="Keyword">private</a>
<a id="2219" class="Keyword">module</a> <a id="Lambek.⊥"></a><a id="2226" href="Categories.Category.Construction.F-Algebras.html#2226" class="Module"></a> <a id="2228" class="Symbol">=</a> <a id="2230" href="Categories.Functor.Algebra.html#487" class="Module">F-Algebra</a> <a id="2240" href="Categories.Object.Initial.html#804" class="Field"></a>
<a id="Lambek.A"></a><a id="2246" href="Categories.Category.Construction.F-Algebras.html#2246" class="Function">A</a> <a id="2248" class="Symbol">=</a> <a id="2250" href="Categories.Functor.Algebra.html#573" class="Function">⊥.A</a>
<a id="Lambek.a"></a><a id="2258" href="Categories.Category.Construction.F-Algebras.html#2258" class="Function">a</a> <a id="2260" class="Symbol">:</a> <a id="2262" href="Categories.Functor.Core.html#432" class="Function">F₀</a> <a id="2265" href="Categories.Category.Construction.F-Algebras.html#2246" class="Function">A</a> <a id="2267" href="Categories.Category.Core.html#575" class="Function Operator"></a> <a id="2269" href="Categories.Category.Construction.F-Algebras.html#2246" class="Function">A</a>
<a id="2275" href="Categories.Category.Construction.F-Algebras.html#2258" class="Function">a</a> <a id="2277" class="Symbol">=</a> <a id="2279" href="Categories.Functor.Algebra.html#587" class="Function">⊥.α</a>
<a id="2288" class="Comment">-- The F-Algebra structure that will make things work</a>
<a id="Lambek.F⊥"></a><a id="2346" href="Categories.Category.Construction.F-Algebras.html#2346" class="Function">F⊥</a> <a id="2349" class="Symbol">:</a> <a id="2351" href="Categories.Functor.Algebra.html#487" class="Record">F-Algebra</a> <a id="2361" href="Categories.Category.Construction.F-Algebras.html#1666" class="Bound">F</a>
<a id="2367" href="Categories.Category.Construction.F-Algebras.html#2346" class="Function">F⊥</a> <a id="2370" class="Symbol">=</a> <a id="2372" href="Categories.Functor.Algebra.html#859" class="Function">iterate</a> <a id="2380" href="Categories.Object.Initial.html#804" class="Field"></a>
<a id="2387" class="Comment">-- By initiality, we get the following morphism</a>
<a id="Lambek.f"></a><a id="2439" href="Categories.Category.Construction.F-Algebras.html#2439" class="Function">f</a> <a id="2441" class="Symbol">:</a> <a id="2443" href="Categories.Functor.Algebra.html#1307" class="Record">F-Algebra-Morphism</a> <a id="2462" href="Categories.Object.Initial.html#804" class="Field"></a> <a id="2464" href="Categories.Category.Construction.F-Algebras.html#2346" class="Function">F⊥</a>
<a id="2471" href="Categories.Category.Construction.F-Algebras.html#2439" class="Function">f</a> <a id="2473" class="Symbol">=</a> <a id="2475" href="Categories.Object.Initial.html#478" class="Function">!</a>
<a id="2482" class="Keyword">module</a> <a id="Lambek.FAM"></a><a id="2489" href="Categories.Category.Construction.F-Algebras.html#2489" class="Module">FAM</a> <a id="2493" class="Symbol">=</a> <a id="2495" href="Categories.Functor.Algebra.html#1307" class="Module">F-Algebra-Morphism</a> <a id="2514" href="Categories.Category.Construction.F-Algebras.html#2439" class="Function">f</a>
<a id="Lambek.i"></a><a id="2521" href="Categories.Category.Construction.F-Algebras.html#2521" class="Function">i</a> <a id="2523" class="Symbol">:</a> <a id="2525" href="Categories.Category.Construction.F-Algebras.html#2246" class="Function">A</a> <a id="2527" href="Categories.Category.Core.html#575" class="Function Operator"></a> <a id="2529" href="Categories.Functor.Core.html#432" class="Function">F₀</a> <a id="2532" href="Categories.Category.Construction.F-Algebras.html#2246" class="Function">A</a>
<a id="2538" href="Categories.Category.Construction.F-Algebras.html#2521" class="Function">i</a> <a id="2540" class="Symbol">=</a> <a id="2542" href="Categories.Functor.Algebra.html#1366" class="Function">FAM.f</a>
<a id="Lambek.a∘f"></a><a id="2553" href="Categories.Category.Construction.F-Algebras.html#2553" class="Function">a∘f</a> <a id="2557" class="Symbol">:</a> <a id="2559" href="Categories.Functor.Algebra.html#1307" class="Record">F-Algebra-Morphism</a> <a id="2578" href="Categories.Object.Initial.html#804" class="Field"></a> <a id="2580" href="Categories.Object.Initial.html#804" class="Field"></a>
<a id="2586" href="Categories.Category.Construction.F-Algebras.html#2553" class="Function">a∘f</a> <a id="2590" class="Symbol">=</a> <a id="2592" class="Keyword">record</a>
<a id="2605" class="Symbol">{</a> <a id="2607" href="Categories.Functor.Algebra.html#1366" class="Field">f</a> <a id="2609" class="Symbol">=</a> <a id="2611" href="Categories.Category.Construction.F-Algebras.html#2258" class="Function">a</a> <a id="2613" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="2615" href="Categories.Category.Construction.F-Algebras.html#2521" class="Function">i</a>
<a id="2623" class="Symbol">;</a> <a id="2625" href="Categories.Functor.Algebra.html#1388" class="Field">commutes</a> <a id="2634" class="Symbol">=</a> <a id="2636" href="Categories.Morphism.Reasoning.Core.html#4290" class="Function">glue</a> <a id="2641" href="Categories.Category.Construction.F-Algebras.html#2757" class="Function">triv</a> <a id="2646" href="Categories.Functor.Algebra.html#1388" class="Function">FAM.commutes</a> <a id="2659" href="Categories.Category.Core.html#3061" class="Function Operator"></a> <a id="2661" href="Categories.Category.Core.html#1706" class="Function">∘-resp-≈ʳ</a> <a id="2671" class="Symbol">(</a><a id="2672" href="Categories.Category.Core.html#3005" class="Function"></a> <a id="2674" href="Categories.Functor.Core.html#565" class="Function">homomorphism</a><a id="2686" class="Symbol">)</a>
<a id="2694" class="Symbol">}</a>
<a id="2702" class="Keyword">where</a>
<a id="2716" class="Keyword">open</a> <a id="2721" href="Categories.Category.Core.html#2462" class="Module">HomReasoning</a> <a id="2734" class="Keyword">using</a> <a id="2740" class="Symbol">(</a><a id="2741" href="Categories.Category.Core.html#3061" class="Function Operator">_○_</a><a id="2744" class="Symbol">;</a> <a id="2746" href="Categories.Category.Core.html#3005" class="Function"></a><a id="2747" class="Symbol">)</a>
<a id="2757" href="Categories.Category.Construction.F-Algebras.html#2757" class="Function">triv</a> <a id="2762" class="Symbol">:</a> <a id="2764" href="Categories.Category.html#884" class="Function">CommutativeSquare</a> <a id="2782" class="Symbol">(</a><a id="2783" href="Categories.Functor.Algebra.html#587" class="Field">α</a> <a id="2785" href="Categories.Category.Construction.F-Algebras.html#2346" class="Function">F⊥</a><a id="2787" class="Symbol">)</a> <a id="2789" class="Symbol">(</a><a id="2790" href="Categories.Functor.Algebra.html#587" class="Field">α</a> <a id="2792" href="Categories.Category.Construction.F-Algebras.html#2346" class="Function">F⊥</a><a id="2794" class="Symbol">)</a> <a id="2796" href="Categories.Category.Construction.F-Algebras.html#2258" class="Function">a</a> <a id="2798" href="Categories.Category.Construction.F-Algebras.html#2258" class="Function">a</a>
<a id="2808" href="Categories.Category.Construction.F-Algebras.html#2757" class="Function">triv</a> <a id="2813" class="Symbol">=</a> <a id="2815" href="Relation.Binary.Structures.html#1596" class="Function">Equiv.refl</a>
<a id="Lambek.lambek"></a><a id="2829" href="Categories.Category.Construction.F-Algebras.html#2829" class="Function">lambek</a> <a id="2836" class="Symbol">:</a> <a id="2838" href="Categories.Category.Construction.F-Algebras.html#2246" class="Function">A</a> <a id="2840" href="Categories.Morphism.html#1958" class="Record Operator"></a> <a id="2842" href="Categories.Functor.Core.html#432" class="Function">F₀</a> <a id="2845" href="Categories.Category.Construction.F-Algebras.html#2246" class="Function">A</a>
<a id="2849" href="Categories.Category.Construction.F-Algebras.html#2829" class="Function">lambek</a> <a id="2856" class="Symbol">=</a> <a id="2858" class="Keyword">record</a>
<a id="2869" class="Symbol">{</a> <a id="2871" href="Categories.Morphism.html#2006" class="Field">from</a> <a id="2876" class="Symbol">=</a> <a id="2878" href="Categories.Category.Construction.F-Algebras.html#2521" class="Function">i</a>
<a id="2884" class="Symbol">;</a> <a id="2886" href="Categories.Morphism.html#2023" class="Field">to</a> <a id="2889" class="Symbol">=</a> <a id="2891" href="Categories.Category.Construction.F-Algebras.html#2258" class="Function">a</a>
<a id="2897" class="Symbol">;</a> <a id="2899" href="Categories.Morphism.html#2040" class="Field">iso</a> <a id="2903" class="Symbol">=</a> <a id="2905" class="Keyword">record</a>
<a id="2918" class="Symbol">{</a> <a id="2920" href="Categories.Morphism.html#1586" class="Field">isoˡ</a> <a id="2925" class="Symbol">=</a> <a id="2927" href="Categories.Object.Initial.html#698" class="Function">⊥-id</a> <a id="2932" href="Categories.Category.Construction.F-Algebras.html#2553" class="Function">a∘f</a>
<a id="2942" class="Symbol">;</a> <a id="2944" href="Categories.Morphism.html#1612" class="Field">isoʳ</a> <a id="2949" class="Symbol">=</a> <a id="2951" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
<a id="2965" href="Categories.Category.Construction.F-Algebras.html#2521" class="Function">i</a> <a id="2967" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="2969" href="Categories.Category.Construction.F-Algebras.html#2258" class="Function">a</a> <a id="2977" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2980" href="Categories.Functor.Algebra.html#1388" class="Field">F-Algebra-Morphism.commutes</a> <a id="3008" href="Categories.Category.Construction.F-Algebras.html#2439" class="Function">f</a> <a id="3010" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="3020" href="Categories.Functor.Core.html#455" class="Function">F₁</a> <a id="3023" href="Categories.Category.Construction.F-Algebras.html#2258" class="Function">a</a> <a id="3025" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="3027" href="Categories.Functor.Core.html#455" class="Function">F₁</a> <a id="3030" href="Categories.Category.Construction.F-Algebras.html#2521" class="Function">i</a> <a id="3032" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function">≈˘⟨</a> <a id="3036" href="Categories.Functor.Core.html#565" class="Function">homomorphism</a> <a id="3049" href="Relation.Binary.Reasoning.Syntax.html#7400" class="Function"></a>
<a id="3059" href="Categories.Functor.Core.html#455" class="Function">F₁</a> <a id="3062" class="Symbol">(</a><a id="3063" href="Categories.Category.Construction.F-Algebras.html#2258" class="Function">a</a> <a id="3065" href="Categories.Category.Core.html#656" class="Function Operator"></a> <a id="3067" href="Categories.Category.Construction.F-Algebras.html#2521" class="Function">i</a><a id="3068" class="Symbol">)</a> <a id="3071" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3074" href="Categories.Functor.Core.html#696" class="Function">F-resp-≈</a> <a id="3083" class="Symbol">(</a><a id="3084" href="Categories.Object.Initial.html#698" class="Function">⊥-id</a> <a id="3089" href="Categories.Category.Construction.F-Algebras.html#2553" class="Function">a∘f</a><a id="3092" class="Symbol">)</a> <a id="3094" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="3104" href="Categories.Functor.Core.html#455" class="Function">F₁</a> <a id="3107" href="Categories.Category.Core.html#630" class="Function">id</a> <a id="3116" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3119" href="Categories.Functor.Core.html#511" class="Function">identity</a> <a id="3128" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function"></a>
<a id="3138" href="Categories.Category.Core.html#630" class="Function">id</a> <a id="3141" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator"></a>
<a id="3149" class="Symbol">}</a>
<a id="3155" class="Symbol">}</a>
<a id="3161" class="Keyword">where</a>
<a id="3173" class="Keyword">open</a> <a id="3178" href="Categories.Category.Core.html#2462" class="Module">HomReasoning</a>
</pre></body></html>