mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
48 lines
17 KiB
HTML
48 lines
17 KiB
HTML
|
<!DOCTYPE HTML>
|
|||
|
<html><head><meta charset="utf-8"><title>Categories.Functor.Construction.LiftSetoids</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">module</a> <a id="44" href="Categories.Functor.Construction.LiftSetoids.html" class="Module">Categories.Functor.Construction.LiftSetoids</a> <a id="88" class="Keyword">where</a>
|
|||
|
|
|||
|
<a id="95" class="Keyword">open</a> <a id="100" class="Keyword">import</a> <a id="107" href="Level.html" class="Module">Level</a> <a id="113" class="Keyword">using</a> <a id="119" class="Symbol">(</a><a id="120" href="Agda.Primitive.html#742" class="Postulate">Level</a><a id="125" class="Symbol">;</a> <a id="127" href="Agda.Primitive.html#961" class="Primitive Operator">_⊔_</a><a id="130" class="Symbol">;</a> <a id="132" href="Level.html#409" class="Record">Lift</a><a id="136" class="Symbol">;</a> <a id="138" href="Level.html#466" class="InductiveConstructor">lift</a><a id="142" class="Symbol">)</a>
|
|||
|
<a id="144" class="Keyword">open</a> <a id="149" class="Keyword">import</a> <a id="156" href="Relation.Binary.Bundles.html" class="Module">Relation.Binary.Bundles</a> <a id="180" class="Keyword">using</a> <a id="186" class="Symbol">(</a><a id="187" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a><a id="193" class="Symbol">)</a>
|
|||
|
<a id="195" class="Keyword">open</a> <a id="200" class="Keyword">import</a> <a id="207" href="Function.Bundles.html" class="Module">Function.Bundles</a> <a id="224" class="Keyword">using</a> <a id="230" class="Symbol">(</a><a id="231" href="Function.Bundles.html#2043" class="Record">Func</a><a id="235" class="Symbol">;</a> <a id="237" href="Function.Bundles.html#15133" class="Function Operator">_⟨$⟩_</a><a id="242" class="Symbol">)</a>
|
|||
|
<a id="244" class="Keyword">open</a> <a id="249" class="Keyword">import</a> <a id="256" href="Function.html" class="Module">Function</a> <a id="265" class="Keyword">using</a> <a id="271" class="Symbol">(</a><a id="272" href="Function.Base.html#1974" class="Function Operator">_$_</a><a id="275" class="Symbol">)</a> <a id="277" class="Keyword">renaming</a> <a id="286" class="Symbol">(</a><a id="287" href="Function.Base.html#704" class="Function">id</a> <a id="290" class="Symbol">to</a> <a id="293" class="Function">idf</a><a id="296" class="Symbol">)</a>
|
|||
|
|
|||
|
<a id="299" class="Keyword">open</a> <a id="304" class="Keyword">import</a> <a id="311" href="Categories.Category.Instance.Setoids.html" class="Module">Categories.Category.Instance.Setoids</a> <a id="348" class="Keyword">using</a> <a id="354" class="Symbol">(</a><a id="355" href="Categories.Category.Instance.Setoids.html#555" class="Function">Setoids</a><a id="362" class="Symbol">)</a>
|
|||
|
<a id="364" class="Keyword">open</a> <a id="369" class="Keyword">import</a> <a id="376" href="Categories.Functor.Core.html" class="Module">Categories.Functor.Core</a> <a id="400" class="Keyword">using</a> <a id="406" class="Symbol">(</a><a id="407" href="Categories.Functor.Core.html#248" class="Record">Functor</a><a id="414" class="Symbol">)</a>
|
|||
|
|
|||
|
<a id="417" class="Keyword">open</a> <a id="422" href="Function.Bundles.html#2043" class="Module">Func</a>
|
|||
|
|
|||
|
<a id="428" class="Keyword">private</a>
|
|||
|
<a id="438" class="Keyword">variable</a>
|
|||
|
<a id="451" href="Categories.Functor.Construction.LiftSetoids.html#451" class="Generalizable">c</a> <a id="453" href="Categories.Functor.Construction.LiftSetoids.html#453" class="Generalizable">ℓ</a> <a id="455" class="Symbol">:</a> <a id="457" href="Agda.Primitive.html#742" class="Postulate">Level</a>
|
|||
|
|
|||
|
<a id="464" class="Comment">-- Use pattern-matching (instead of explicit calls to lower) to minimize the</a>
|
|||
|
<a id="541" class="Comment">-- number of needed parens, and also make it syntactically apparent that</a>
|
|||
|
<a id="614" class="Comment">-- this is indeed just a Lift.</a>
|
|||
|
<a id="LiftedSetoid"></a><a id="645" href="Categories.Functor.Construction.LiftSetoids.html#645" class="Function">LiftedSetoid</a> <a id="658" class="Symbol">:</a> <a id="660" class="Symbol">∀</a> <a id="662" href="Categories.Functor.Construction.LiftSetoids.html#662" class="Bound">c′</a> <a id="665" href="Categories.Functor.Construction.LiftSetoids.html#665" class="Bound">ℓ′</a> <a id="668" class="Symbol">→</a> <a id="670" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="677" href="Categories.Functor.Construction.LiftSetoids.html#451" class="Generalizable">c</a> <a id="679" href="Categories.Functor.Construction.LiftSetoids.html#453" class="Generalizable">ℓ</a> <a id="681" class="Symbol">→</a> <a id="683" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="690" class="Symbol">(</a><a id="691" href="Categories.Functor.Construction.LiftSetoids.html#451" class="Generalizable">c</a> <a id="693" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="695" href="Categories.Functor.Construction.LiftSetoids.html#662" class="Bound">c′</a><a id="697" class="Symbol">)</a> <a id="699" class="Symbol">(</a><a id="700" href="Categories.Functor.Construction.LiftSetoids.html#453" class="Generalizable">ℓ</a> <a id="702" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="704" href="Categories.Functor.Construction.LiftSetoids.html#665" class="Bound">ℓ′</a><a id="706" class="Symbol">)</a>
|
|||
|
<a id="708" href="Categories.Functor.Construction.LiftSetoids.html#645" class="Function">LiftedSetoid</a> <a id="721" href="Categories.Functor.Construction.LiftSetoids.html#721" class="Bound">c′</a> <a id="724" href="Categories.Functor.Construction.LiftSetoids.html#724" class="Bound">ℓ′</a> <a id="727" href="Categories.Functor.Construction.LiftSetoids.html#727" class="Bound">S</a> <a id="729" class="Symbol">=</a> <a id="731" class="Keyword">record</a>
|
|||
|
<a id="740" class="Symbol">{</a> <a id="742" href="Relation.Binary.Bundles.html#1143" class="Field">Carrier</a> <a id="756" class="Symbol">=</a> <a id="758" href="Level.html#409" class="Record">Lift</a> <a id="763" href="Categories.Functor.Construction.LiftSetoids.html#721" class="Bound">c′</a> <a id="766" href="Relation.Binary.Bundles.html#1143" class="Field">Carrier</a>
|
|||
|
<a id="776" class="Symbol">;</a> <a id="778" href="Relation.Binary.Bundles.html#1169" class="Field Operator">_≈_</a> <a id="792" class="Symbol">=</a> <a id="794" class="Symbol">λ</a> <a id="796" class="Keyword">where</a> <a id="802" class="Symbol">(</a><a id="803" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="808" href="Categories.Functor.Construction.LiftSetoids.html#808" class="Bound">x</a><a id="809" class="Symbol">)</a> <a id="811" class="Symbol">(</a><a id="812" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="817" href="Categories.Functor.Construction.LiftSetoids.html#817" class="Bound">y</a><a id="818" class="Symbol">)</a> <a id="820" class="Symbol">→</a> <a id="822" href="Level.html#409" class="Record">Lift</a> <a id="827" href="Categories.Functor.Construction.LiftSetoids.html#724" class="Bound">ℓ′</a> <a id="830" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="832" href="Categories.Functor.Construction.LiftSetoids.html#808" class="Bound">x</a> <a id="834" href="Relation.Binary.Bundles.html#1169" class="Field Operator">≈</a> <a id="836" href="Categories.Functor.Construction.LiftSetoids.html#817" class="Bound">y</a>
|
|||
|
<a id="840" class="Symbol">;</a> <a id="842" href="Relation.Binary.Bundles.html#1203" class="Field">isEquivalence</a> <a id="856" class="Symbol">=</a> <a id="858" class="Keyword">record</a>
|
|||
|
<a id="869" class="Symbol">{</a> <a id="871" href="Relation.Binary.Structures.html#1596" class="Field">refl</a> <a id="877" class="Symbol">=</a> <a id="879" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="884" href="Relation.Binary.Structures.html#1596" class="Function">refl</a>
|
|||
|
<a id="893" class="Symbol">;</a> <a id="895" href="Relation.Binary.Structures.html#1622" class="Field">sym</a> <a id="901" class="Symbol">=</a> <a id="903" class="Symbol">λ</a> <a id="905" class="Keyword">where</a> <a id="911" class="Symbol">(</a><a id="912" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="917" href="Categories.Functor.Construction.LiftSetoids.html#917" class="Bound">eq</a><a id="919" class="Symbol">)</a> <a id="932" class="Symbol">→</a> <a id="934" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="939" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="941" href="Relation.Binary.Structures.html#1200" class="Function">sym</a> <a id="945" href="Categories.Functor.Construction.LiftSetoids.html#917" class="Bound">eq</a>
|
|||
|
<a id="952" class="Symbol">;</a> <a id="954" href="Relation.Binary.Structures.html#1648" class="Field">trans</a> <a id="960" class="Symbol">=</a> <a id="962" class="Symbol">λ</a> <a id="964" class="Keyword">where</a> <a id="970" class="Symbol">(</a><a id="971" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="976" href="Categories.Functor.Construction.LiftSetoids.html#976" class="Bound">eq</a><a id="978" class="Symbol">)</a> <a id="980" class="Symbol">(</a><a id="981" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="986" href="Categories.Functor.Construction.LiftSetoids.html#986" class="Bound">eq′</a><a id="989" class="Symbol">)</a> <a id="991" class="Symbol">→</a> <a id="993" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="998" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="1000" href="Relation.Binary.Structures.html#1226" class="Function">trans</a> <a id="1006" href="Categories.Functor.Construction.LiftSetoids.html#976" class="Bound">eq</a> <a id="1009" href="Categories.Functor.Construction.LiftSetoids.html#986" class="Bound">eq′</a>
|
|||
|
<a id="1017" class="Symbol">}</a>
|
|||
|
<a id="1021" class="Symbol">}</a>
|
|||
|
<a id="1025" class="Keyword">where</a> <a id="1031" class="Keyword">open</a> <a id="1036" href="Relation.Binary.Bundles.html#1080" class="Module">Setoid</a> <a id="1043" href="Categories.Functor.Construction.LiftSetoids.html#727" class="Bound">S</a>
|
|||
|
|
|||
|
<a id="LiftSetoids"></a><a id="1046" href="Categories.Functor.Construction.LiftSetoids.html#1046" class="Function">LiftSetoids</a> <a id="1058" class="Symbol">:</a> <a id="1060" class="Symbol">∀</a> <a id="1062" href="Categories.Functor.Construction.LiftSetoids.html#1062" class="Bound">c′</a> <a id="1065" href="Categories.Functor.Construction.LiftSetoids.html#1065" class="Bound">ℓ′</a> <a id="1068" class="Symbol">→</a> <a id="1070" href="Categories.Functor.Core.html#248" class="Record">Functor</a> <a id="1078" class="Symbol">(</a><a id="1079" href="Categories.Category.Instance.Setoids.html#555" class="Function">Setoids</a> <a id="1087" href="Categories.Functor.Construction.LiftSetoids.html#451" class="Generalizable">c</a> <a id="1089" href="Categories.Functor.Construction.LiftSetoids.html#453" class="Generalizable">ℓ</a><a id="1090" class="Symbol">)</a> <a id="1092" class="Symbol">(</a><a id="1093" href="Categories.Category.Instance.Setoids.html#555" class="Function">Setoids</a> <a id="1101" class="Symbol">(</a><a id="1102" href="Categories.Functor.Construction.LiftSetoids.html#451" class="Generalizable">c</a> <a id="1104" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1106" href="Categories.Functor.Construction.LiftSetoids.html#1062" class="Bound">c′</a><a id="1108" class="Symbol">)</a> <a id="1110" class="Symbol">(</a><a id="1111" href="Categories.Functor.Construction.LiftSetoids.html#453" class="Generalizable">ℓ</a> <a id="1113" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1115" href="Categories.Functor.Construction.LiftSetoids.html#1065" class="Bound">ℓ′</a><a id="1117" class="Symbol">))</a>
|
|||
|
<a id="1120" href="Categories.Functor.Construction.LiftSetoids.html#1046" class="Function">LiftSetoids</a> <a id="1132" href="Categories.Functor.Construction.LiftSetoids.html#1132" class="Bound">c′</a> <a id="1135" href="Categories.Functor.Construction.LiftSetoids.html#1135" class="Bound">ℓ′</a> <a id="1138" class="Symbol">=</a> <a id="1140" class="Keyword">record</a>
|
|||
|
<a id="1149" class="Symbol">{</a> <a id="1151" href="Categories.Functor.Core.html#432" class="Field">F₀</a> <a id="1164" class="Symbol">=</a> <a id="1166" href="Categories.Functor.Construction.LiftSetoids.html#645" class="Function">LiftedSetoid</a> <a id="1179" href="Categories.Functor.Construction.LiftSetoids.html#1132" class="Bound">c′</a> <a id="1182" href="Categories.Functor.Construction.LiftSetoids.html#1135" class="Bound">ℓ′</a>
|
|||
|
<a id="1187" class="Symbol">;</a> <a id="1189" href="Categories.Functor.Core.html#455" class="Field">F₁</a> <a id="1202" class="Symbol">=</a> <a id="1204" class="Symbol">λ</a> <a id="1206" href="Categories.Functor.Construction.LiftSetoids.html#1206" class="Bound">f</a> <a id="1208" class="Symbol">→</a> <a id="1210" class="Keyword">record</a>
|
|||
|
<a id="1221" class="Symbol">{</a> <a id="1223" href="Function.Bundles.html#2094" class="Field">to</a> <a id="1229" class="Symbol">=</a> <a id="1231" class="Symbol">λ</a> <a id="1233" class="Keyword">where</a> <a id="1239" class="Symbol">(</a><a id="1240" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="1245" href="Categories.Functor.Construction.LiftSetoids.html#1245" class="Bound">x</a><a id="1246" class="Symbol">)</a> <a id="1249" class="Symbol">→</a> <a id="1251" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="1256" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="1258" href="Categories.Functor.Construction.LiftSetoids.html#1206" class="Bound">f</a> <a id="1260" href="Function.Bundles.html#15133" class="Function Operator">⟨$⟩</a> <a id="1264" href="Categories.Functor.Construction.LiftSetoids.html#1245" class="Bound">x</a>
|
|||
|
<a id="1270" class="Symbol">;</a> <a id="1272" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="1278" class="Symbol">=</a> <a id="1280" class="Symbol">λ</a> <a id="1282" class="Keyword">where</a> <a id="1288" class="Symbol">(</a><a id="1289" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="1294" href="Categories.Functor.Construction.LiftSetoids.html#1294" class="Bound">eq</a><a id="1296" class="Symbol">)</a> <a id="1298" class="Symbol">→</a> <a id="1300" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="1305" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="1307" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="1312" href="Categories.Functor.Construction.LiftSetoids.html#1206" class="Bound">f</a> <a id="1314" href="Categories.Functor.Construction.LiftSetoids.html#1294" class="Bound">eq</a>
|
|||
|
<a id="1321" class="Symbol">}</a>
|
|||
|
<a id="1325" class="Symbol">;</a> <a id="1327" href="Categories.Functor.Core.html#511" class="Field">identity</a> <a id="1340" class="Symbol">=</a> <a id="1342" class="Symbol">λ</a> <a id="1344" class="Symbol">{</a><a id="1345" href="Categories.Functor.Construction.LiftSetoids.html#1345" class="Bound">A</a><a id="1346" class="Symbol">}</a> <a id="1348" class="Symbol">→</a> <a id="1350" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="1355" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="1357" href="Relation.Binary.Structures.html#1596" class="Function">refl</a> <a id="1362" href="Categories.Functor.Construction.LiftSetoids.html#1345" class="Bound">A</a>
|
|||
|
<a id="1366" class="Symbol">;</a> <a id="1368" href="Categories.Functor.Core.html#565" class="Field">homomorphism</a> <a id="1381" class="Symbol">=</a> <a id="1383" class="Symbol">λ</a> <a id="1385" class="Symbol">{</a><a id="1386" href="Categories.Functor.Construction.LiftSetoids.html#1386" class="Bound">_</a><a id="1387" class="Symbol">}</a> <a id="1389" class="Symbol">{</a><a id="1390" href="Categories.Functor.Construction.LiftSetoids.html#1390" class="Bound">_</a><a id="1391" class="Symbol">}</a> <a id="1393" class="Symbol">{</a><a id="1394" href="Categories.Functor.Construction.LiftSetoids.html#1394" class="Bound">Z</a><a id="1395" class="Symbol">}</a> <a id="1397" class="Symbol">→</a> <a id="1399" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="1404" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="1406" href="Relation.Binary.Structures.html#1596" class="Function">refl</a> <a id="1411" href="Categories.Functor.Construction.LiftSetoids.html#1394" class="Bound">Z</a>
|
|||
|
<a id="1415" class="Symbol">;</a> <a id="1417" href="Categories.Functor.Core.html#696" class="Field">F-resp-≈</a> <a id="1430" class="Symbol">=</a> <a id="1432" class="Symbol">λ</a> <a id="1434" href="Categories.Functor.Construction.LiftSetoids.html#1434" class="Bound">f≈g</a> <a id="1438" class="Symbol">→</a> <a id="1440" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="1445" href="Categories.Functor.Construction.LiftSetoids.html#1434" class="Bound">f≈g</a>
|
|||
|
<a id="1451" class="Symbol">}</a>
|
|||
|
<a id="1455" class="Keyword">where</a>
|
|||
|
<a id="1465" class="Keyword">open</a> <a id="1470" href="Relation.Binary.Bundles.html#1080" class="Module">Setoid</a>
|
|||
|
</pre></body></html>
|