bsc-leon-vatthauer/agda/bsc-thesis/Algebra.Construct.NaturalChoice.Base.html
2024-02-09 17:53:52 +01:00

63 lines
No EOL
17 KiB
HTML

<!DOCTYPE HTML>
<html><head><meta charset="utf-8"><title>Algebra.Construct.NaturalChoice.Base</title><link rel="stylesheet" href="Agda.css"></head><body><pre class="Agda"><a id="1" class="Comment">------------------------------------------------------------------------</a>
<a id="74" class="Comment">-- The Agda standard library</a>
<a id="103" class="Comment">--</a>
<a id="106" class="Comment">-- Basic definition of an operator that computes the min/max value</a>
<a id="173" class="Comment">-- with respect to a total preorder.</a>
<a id="210" class="Comment">------------------------------------------------------------------------</a>
<a id="284" class="Symbol">{-#</a> <a id="288" class="Keyword">OPTIONS</a> <a id="296" class="Pragma">--cubical-compatible</a> <a id="317" class="Pragma">--safe</a> <a id="324" class="Symbol">#-}</a>
<a id="329" class="Keyword">open</a> <a id="334" class="Keyword">import</a> <a id="341" href="Algebra.Core.html" class="Module">Algebra.Core</a>
<a id="354" class="Keyword">open</a> <a id="359" class="Keyword">import</a> <a id="366" href="Level.html" class="Module">Level</a> <a id="372" class="Symbol">as</a> <a id="375" class="Module">L</a> <a id="377" class="Keyword">hiding</a> <a id="384" class="Symbol">(</a><a id="385" href="Agda.Primitive.html#961" class="Primitive Operator">_⊔_</a><a id="388" class="Symbol">)</a>
<a id="390" class="Keyword">open</a> <a id="395" class="Keyword">import</a> <a id="402" href="Function.Base.html" class="Module">Function.Base</a> <a id="416" class="Keyword">using</a> <a id="422" class="Symbol">(</a><a id="423" href="Function.Base.html#1638" class="Function">flip</a><a id="427" class="Symbol">)</a>
<a id="429" class="Keyword">open</a> <a id="434" class="Keyword">import</a> <a id="441" href="Relation.Binary.Bundles.html" class="Module">Relation.Binary.Bundles</a> <a id="465" class="Keyword">using</a> <a id="471" class="Symbol">(</a><a id="472" href="Relation.Binary.Bundles.html#2842" class="Record">TotalPreorder</a><a id="485" class="Symbol">)</a>
<a id="487" class="Keyword">open</a> <a id="492" class="Keyword">import</a> <a id="499" href="Relation.Binary.Construct.Flip.EqAndOrd.html" class="Module">Relation.Binary.Construct.Flip.EqAndOrd</a> <a id="539" class="Keyword">using</a> <a id="545" class="Symbol">()</a>
<a id="550" class="Keyword">renaming</a> <a id="559" class="Symbol">(</a><a id="560" href="Relation.Binary.Construct.Flip.EqAndOrd.html#5603" class="Function">totalPreorder</a> <a id="574" class="Symbol">to</a> <a id="577" class="Function">flipOrder</a><a id="586" class="Symbol">)</a>
<a id="588" class="Keyword">import</a> <a id="595" href="Relation.Binary.Properties.TotalOrder.html" class="Module">Relation.Binary.Properties.TotalOrder</a> <a id="633" class="Symbol">as</a> <a id="636" class="Module">TotalOrderProperties</a>
<a id="658" class="Keyword">module</a> <a id="665" href="Algebra.Construct.NaturalChoice.Base.html" class="Module">Algebra.Construct.NaturalChoice.Base</a> <a id="702" class="Keyword">where</a>
<a id="709" class="Keyword">private</a>
<a id="719" class="Keyword">variable</a>
<a id="732" href="Algebra.Construct.NaturalChoice.Base.html#732" class="Generalizable">a</a> <a id="734" href="Algebra.Construct.NaturalChoice.Base.html#734" class="Generalizable">ℓ₁</a> <a id="737" href="Algebra.Construct.NaturalChoice.Base.html#737" class="Generalizable">ℓ₂</a> <a id="740" class="Symbol">:</a> <a id="742" href="Agda.Primitive.html#742" class="Postulate">Level</a>
<a id="752" href="Algebra.Construct.NaturalChoice.Base.html#752" class="Generalizable">O</a> <a id="754" class="Symbol">:</a> <a id="756" href="Relation.Binary.Bundles.html#2842" class="Record">TotalPreorder</a> <a id="770" href="Algebra.Construct.NaturalChoice.Base.html#732" class="Generalizable">a</a> <a id="772" href="Algebra.Construct.NaturalChoice.Base.html#734" class="Generalizable">ℓ₁</a> <a id="775" href="Algebra.Construct.NaturalChoice.Base.html#737" class="Generalizable">ℓ₂</a>
<a id="779" class="Comment">------------------------------------------------------------------------</a>
<a id="852" class="Comment">-- Definition</a>
<a id="867" class="Keyword">module</a> <a id="874" href="Algebra.Construct.NaturalChoice.Base.html#874" class="Module">_</a> <a id="876" class="Symbol">(</a><a id="877" href="Algebra.Construct.NaturalChoice.Base.html#877" class="Bound">O</a> <a id="879" class="Symbol">:</a> <a id="881" href="Relation.Binary.Bundles.html#2842" class="Record">TotalPreorder</a> <a id="895" href="Algebra.Construct.NaturalChoice.Base.html#732" class="Generalizable">a</a> <a id="897" href="Algebra.Construct.NaturalChoice.Base.html#734" class="Generalizable">ℓ₁</a> <a id="900" href="Algebra.Construct.NaturalChoice.Base.html#737" class="Generalizable">ℓ₂</a><a id="902" class="Symbol">)</a> <a id="904" class="Keyword">where</a>
<a id="912" class="Keyword">open</a> <a id="917" href="Relation.Binary.Bundles.html#2842" class="Module">TotalPreorder</a> <a id="931" href="Algebra.Construct.NaturalChoice.Base.html#877" class="Bound">O</a> <a id="933" class="Keyword">renaming</a> <a id="942" class="Symbol">(</a><a id="943" href="Relation.Binary.Bundles.html#3020" class="Field Operator">_≲_</a> <a id="947" class="Symbol">to</a> <a id="950" class="Field Operator">_≤_</a><a id="953" class="Symbol">)</a>
<a id="957" class="Keyword">private</a> <a id="965" href="Algebra.Construct.NaturalChoice.Base.html#965" class="Function Operator">_≥_</a> <a id="969" class="Symbol">=</a> <a id="971" href="Function.Base.html#1638" class="Function">flip</a> <a id="976" href="Algebra.Construct.NaturalChoice.Base.html#950" class="Field Operator">_≤_</a>
<a id="983" class="Keyword">record</a> <a id="990" href="Algebra.Construct.NaturalChoice.Base.html#990" class="Record">MinOperator</a> <a id="1002" class="Symbol">:</a> <a id="1004" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1008" class="Symbol">(</a><a id="1009" href="Algebra.Construct.NaturalChoice.Base.html#895" class="Bound">a</a> <a id="1011" href="Agda.Primitive.html#961" class="Primitive Operator">L.⊔</a> <a id="1015" href="Algebra.Construct.NaturalChoice.Base.html#897" class="Bound">ℓ₁</a> <a id="1018" href="Agda.Primitive.html#961" class="Primitive Operator">L.⊔</a> <a id="1022" href="Algebra.Construct.NaturalChoice.Base.html#900" class="Bound">ℓ₂</a><a id="1024" class="Symbol">)</a> <a id="1026" class="Keyword">where</a>
<a id="1036" class="Keyword">infixl</a> <a id="1043" class="Number">7</a> <a id="1045" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Field Operator">_⊓_</a>
<a id="1053" class="Keyword">field</a>
<a id="1065" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Field Operator">_⊓_</a> <a id="1075" class="Symbol">:</a> <a id="1077" href="Algebra.Core.html#527" class="Function">Op₂</a> <a id="1081" href="Relation.Binary.Bundles.html#2926" class="Field">Carrier</a>
<a id="1095" href="Algebra.Construct.NaturalChoice.Base.html#1095" class="Field">x≤y⇒x⊓y≈x</a> <a id="1105" class="Symbol">:</a> <a id="1107" class="Symbol"></a> <a id="1109" class="Symbol">{</a><a id="1110" href="Algebra.Construct.NaturalChoice.Base.html#1110" class="Bound">x</a> <a id="1112" href="Algebra.Construct.NaturalChoice.Base.html#1112" class="Bound">y</a><a id="1113" class="Symbol">}</a> <a id="1115" class="Symbol"></a> <a id="1117" href="Algebra.Construct.NaturalChoice.Base.html#1110" class="Bound">x</a> <a id="1119" href="Algebra.Construct.NaturalChoice.Base.html#950" class="Field Operator"></a> <a id="1121" href="Algebra.Construct.NaturalChoice.Base.html#1112" class="Bound">y</a> <a id="1123" class="Symbol"></a> <a id="1125" href="Algebra.Construct.NaturalChoice.Base.html#1110" class="Bound">x</a> <a id="1127" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Field Operator"></a> <a id="1129" href="Algebra.Construct.NaturalChoice.Base.html#1112" class="Bound">y</a> <a id="1131" href="Relation.Binary.Bundles.html#2954" class="Field Operator"></a> <a id="1133" href="Algebra.Construct.NaturalChoice.Base.html#1110" class="Bound">x</a>
<a id="1141" href="Algebra.Construct.NaturalChoice.Base.html#1141" class="Field">x≥y⇒x⊓y≈y</a> <a id="1151" class="Symbol">:</a> <a id="1153" class="Symbol"></a> <a id="1155" class="Symbol">{</a><a id="1156" href="Algebra.Construct.NaturalChoice.Base.html#1156" class="Bound">x</a> <a id="1158" href="Algebra.Construct.NaturalChoice.Base.html#1158" class="Bound">y</a><a id="1159" class="Symbol">}</a> <a id="1161" class="Symbol"></a> <a id="1163" href="Algebra.Construct.NaturalChoice.Base.html#1156" class="Bound">x</a> <a id="1165" href="Algebra.Construct.NaturalChoice.Base.html#965" class="Function Operator"></a> <a id="1167" href="Algebra.Construct.NaturalChoice.Base.html#1158" class="Bound">y</a> <a id="1169" class="Symbol"></a> <a id="1171" href="Algebra.Construct.NaturalChoice.Base.html#1156" class="Bound">x</a> <a id="1173" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Field Operator"></a> <a id="1175" href="Algebra.Construct.NaturalChoice.Base.html#1158" class="Bound">y</a> <a id="1177" href="Relation.Binary.Bundles.html#2954" class="Field Operator"></a> <a id="1179" href="Algebra.Construct.NaturalChoice.Base.html#1158" class="Bound">y</a>
<a id="1184" class="Keyword">record</a> <a id="1191" href="Algebra.Construct.NaturalChoice.Base.html#1191" class="Record">MaxOperator</a> <a id="1203" class="Symbol">:</a> <a id="1205" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1209" class="Symbol">(</a><a id="1210" href="Algebra.Construct.NaturalChoice.Base.html#895" class="Bound">a</a> <a id="1212" href="Agda.Primitive.html#961" class="Primitive Operator">L.⊔</a> <a id="1216" href="Algebra.Construct.NaturalChoice.Base.html#897" class="Bound">ℓ₁</a> <a id="1219" href="Agda.Primitive.html#961" class="Primitive Operator">L.⊔</a> <a id="1223" href="Algebra.Construct.NaturalChoice.Base.html#900" class="Bound">ℓ₂</a><a id="1225" class="Symbol">)</a> <a id="1227" class="Keyword">where</a>
<a id="1237" class="Keyword">infixl</a> <a id="1244" class="Number">6</a> <a id="1246" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a>
<a id="1254" class="Keyword">field</a>
<a id="1266" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a> <a id="1276" class="Symbol">:</a> <a id="1278" href="Algebra.Core.html#527" class="Function">Op₂</a> <a id="1282" href="Relation.Binary.Bundles.html#2926" class="Field">Carrier</a>
<a id="1296" href="Algebra.Construct.NaturalChoice.Base.html#1296" class="Field">x≤y⇒x⊔y≈y</a> <a id="1306" class="Symbol">:</a> <a id="1308" class="Symbol"></a> <a id="1310" class="Symbol">{</a><a id="1311" href="Algebra.Construct.NaturalChoice.Base.html#1311" class="Bound">x</a> <a id="1313" href="Algebra.Construct.NaturalChoice.Base.html#1313" class="Bound">y</a><a id="1314" class="Symbol">}</a> <a id="1316" class="Symbol"></a> <a id="1318" href="Algebra.Construct.NaturalChoice.Base.html#1311" class="Bound">x</a> <a id="1320" href="Algebra.Construct.NaturalChoice.Base.html#950" class="Field Operator"></a> <a id="1322" href="Algebra.Construct.NaturalChoice.Base.html#1313" class="Bound">y</a> <a id="1324" class="Symbol"></a> <a id="1326" href="Algebra.Construct.NaturalChoice.Base.html#1311" class="Bound">x</a> <a id="1328" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator"></a> <a id="1330" href="Algebra.Construct.NaturalChoice.Base.html#1313" class="Bound">y</a> <a id="1332" href="Relation.Binary.Bundles.html#2954" class="Field Operator"></a> <a id="1334" href="Algebra.Construct.NaturalChoice.Base.html#1313" class="Bound">y</a>
<a id="1342" href="Algebra.Construct.NaturalChoice.Base.html#1342" class="Field">x≥y⇒x⊔y≈x</a> <a id="1352" class="Symbol">:</a> <a id="1354" class="Symbol"></a> <a id="1356" class="Symbol">{</a><a id="1357" href="Algebra.Construct.NaturalChoice.Base.html#1357" class="Bound">x</a> <a id="1359" href="Algebra.Construct.NaturalChoice.Base.html#1359" class="Bound">y</a><a id="1360" class="Symbol">}</a> <a id="1362" class="Symbol"></a> <a id="1364" href="Algebra.Construct.NaturalChoice.Base.html#1357" class="Bound">x</a> <a id="1366" href="Algebra.Construct.NaturalChoice.Base.html#965" class="Function Operator"></a> <a id="1368" href="Algebra.Construct.NaturalChoice.Base.html#1359" class="Bound">y</a> <a id="1370" class="Symbol"></a> <a id="1372" href="Algebra.Construct.NaturalChoice.Base.html#1357" class="Bound">x</a> <a id="1374" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator"></a> <a id="1376" href="Algebra.Construct.NaturalChoice.Base.html#1359" class="Bound">y</a> <a id="1378" href="Relation.Binary.Bundles.html#2954" class="Field Operator"></a> <a id="1380" href="Algebra.Construct.NaturalChoice.Base.html#1357" class="Bound">x</a>
<a id="1383" class="Comment">------------------------------------------------------------------------</a>
<a id="1456" class="Comment">-- Properties</a>
<a id="MinOp⇒MaxOp"></a><a id="1471" href="Algebra.Construct.NaturalChoice.Base.html#1471" class="Function">MinOp⇒MaxOp</a> <a id="1483" class="Symbol">:</a> <a id="1485" href="Algebra.Construct.NaturalChoice.Base.html#990" class="Record">MinOperator</a> <a id="1497" href="Algebra.Construct.NaturalChoice.Base.html#752" class="Generalizable">O</a> <a id="1499" class="Symbol"></a> <a id="1501" href="Algebra.Construct.NaturalChoice.Base.html#1191" class="Record">MaxOperator</a> <a id="1513" class="Symbol">(</a><a id="1514" href="Algebra.Construct.NaturalChoice.Base.html#577" class="Function">flipOrder</a> <a id="1524" href="Algebra.Construct.NaturalChoice.Base.html#752" class="Generalizable">O</a><a id="1525" class="Symbol">)</a>
<a id="1527" href="Algebra.Construct.NaturalChoice.Base.html#1471" class="Function">MinOp⇒MaxOp</a> <a id="1539" href="Algebra.Construct.NaturalChoice.Base.html#1539" class="Bound">minOp</a> <a id="1545" class="Symbol">=</a> <a id="1547" class="Keyword">record</a>
<a id="1556" class="Symbol">{</a> <a id="1558" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a> <a id="1568" class="Symbol">=</a> <a id="1570" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Field Operator">_⊓_</a>
<a id="1576" class="Symbol">;</a> <a id="1578" href="Algebra.Construct.NaturalChoice.Base.html#1296" class="Field">x≤y⇒x⊔y≈y</a> <a id="1588" class="Symbol">=</a> <a id="1590" href="Algebra.Construct.NaturalChoice.Base.html#1141" class="Field">x≥y⇒x⊓y≈y</a>
<a id="1602" class="Symbol">;</a> <a id="1604" href="Algebra.Construct.NaturalChoice.Base.html#1342" class="Field">x≥y⇒x⊔y≈x</a> <a id="1614" class="Symbol">=</a> <a id="1616" href="Algebra.Construct.NaturalChoice.Base.html#1095" class="Field">x≤y⇒x⊓y≈x</a>
<a id="1628" class="Symbol">}</a> <a id="1630" class="Keyword">where</a> <a id="1636" class="Keyword">open</a> <a id="1641" href="Algebra.Construct.NaturalChoice.Base.html#990" class="Module">MinOperator</a> <a id="1653" href="Algebra.Construct.NaturalChoice.Base.html#1539" class="Bound">minOp</a>
<a id="MaxOp⇒MinOp"></a><a id="1660" href="Algebra.Construct.NaturalChoice.Base.html#1660" class="Function">MaxOp⇒MinOp</a> <a id="1672" class="Symbol">:</a> <a id="1674" href="Algebra.Construct.NaturalChoice.Base.html#1191" class="Record">MaxOperator</a> <a id="1686" href="Algebra.Construct.NaturalChoice.Base.html#752" class="Generalizable">O</a> <a id="1688" class="Symbol"></a> <a id="1690" href="Algebra.Construct.NaturalChoice.Base.html#990" class="Record">MinOperator</a> <a id="1702" class="Symbol">(</a><a id="1703" href="Algebra.Construct.NaturalChoice.Base.html#577" class="Function">flipOrder</a> <a id="1713" href="Algebra.Construct.NaturalChoice.Base.html#752" class="Generalizable">O</a><a id="1714" class="Symbol">)</a>
<a id="1716" href="Algebra.Construct.NaturalChoice.Base.html#1660" class="Function">MaxOp⇒MinOp</a> <a id="1728" href="Algebra.Construct.NaturalChoice.Base.html#1728" class="Bound">maxOp</a> <a id="1734" class="Symbol">=</a> <a id="1736" class="Keyword">record</a>
<a id="1745" class="Symbol">{</a> <a id="1747" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Field Operator">_⊓_</a> <a id="1757" class="Symbol">=</a> <a id="1759" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a>
<a id="1765" class="Symbol">;</a> <a id="1767" href="Algebra.Construct.NaturalChoice.Base.html#1095" class="Field">x≤y⇒x⊓y≈x</a> <a id="1777" class="Symbol">=</a> <a id="1779" href="Algebra.Construct.NaturalChoice.Base.html#1342" class="Field">x≥y⇒x⊔y≈x</a>
<a id="1791" class="Symbol">;</a> <a id="1793" href="Algebra.Construct.NaturalChoice.Base.html#1141" class="Field">x≥y⇒x⊓y≈y</a> <a id="1803" class="Symbol">=</a> <a id="1805" href="Algebra.Construct.NaturalChoice.Base.html#1296" class="Field">x≤y⇒x⊔y≈y</a>
<a id="1817" class="Symbol">}</a> <a id="1819" class="Keyword">where</a> <a id="1825" class="Keyword">open</a> <a id="1830" href="Algebra.Construct.NaturalChoice.Base.html#1191" class="Module">MaxOperator</a> <a id="1842" href="Algebra.Construct.NaturalChoice.Base.html#1728" class="Bound">maxOp</a>
</pre></body></html>