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

200 lines
No EOL
40 KiB
HTML
Raw Blame History

This file contains ambiguous Unicode characters

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

<!DOCTYPE HTML>
<html><head><meta charset="utf-8"><title>Algebra.Lattice.Structures</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">-- Some lattice-like structures defined by properties of _∧_ and __</a>
<a id="175" class="Comment">-- (not packed up with sets, operations, etc.)</a>
<a id="222" class="Comment">--</a>
<a id="225" class="Comment">-- For lattices defined via an order relation, see</a>
<a id="276" class="Comment">-- Relation.Binary.Lattice.</a>
<a id="304" class="Comment">------------------------------------------------------------------------</a>
<a id="378" class="Comment">-- The contents of this module should be accessed via `Algebra.Lattice`,</a>
<a id="451" class="Comment">-- unless you want to parameterise it via the equality relation.</a>
<a id="517" class="Symbol">{-#</a> <a id="521" class="Keyword">OPTIONS</a> <a id="529" class="Pragma">--cubical-compatible</a> <a id="550" class="Pragma">--safe</a> <a id="557" class="Symbol">#-}</a>
<a id="562" class="Keyword">open</a> <a id="567" class="Keyword">import</a> <a id="574" href="Algebra.Core.html" class="Module">Algebra.Core</a>
<a id="587" class="Keyword">open</a> <a id="592" class="Keyword">import</a> <a id="599" href="Data.Product.Base.html" class="Module">Data.Product.Base</a> <a id="617" class="Keyword">using</a> <a id="623" class="Symbol">(</a><a id="624" href="Data.Product.Base.html#636" class="Field">proj₁</a><a id="629" class="Symbol">;</a> <a id="631" href="Data.Product.Base.html#650" class="Field">proj₂</a><a id="636" class="Symbol">)</a>
<a id="638" class="Keyword">open</a> <a id="643" class="Keyword">import</a> <a id="650" href="Level.html" class="Module">Level</a> <a id="656" class="Keyword">using</a> <a id="662" class="Symbol">(</a><a id="663" href="Agda.Primitive.html#961" class="Primitive Operator">_⊔_</a><a id="666" class="Symbol">)</a>
<a id="668" class="Keyword">open</a> <a id="673" class="Keyword">import</a> <a id="680" href="Relation.Binary.Core.html" class="Module">Relation.Binary.Core</a> <a id="701" class="Keyword">using</a> <a id="707" class="Symbol">(</a><a id="708" href="Relation.Binary.Core.html#896" class="Function">Rel</a><a id="711" class="Symbol">)</a>
<a id="713" class="Keyword">open</a> <a id="718" class="Keyword">import</a> <a id="725" href="Relation.Binary.Bundles.html" class="Module">Relation.Binary.Bundles</a> <a id="749" class="Keyword">using</a> <a id="755" class="Symbol">(</a><a id="756" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a><a id="762" class="Symbol">)</a>
<a id="764" class="Keyword">open</a> <a id="769" class="Keyword">import</a> <a id="776" href="Relation.Binary.Structures.html" class="Module">Relation.Binary.Structures</a> <a id="803" class="Keyword">using</a> <a id="809" class="Symbol">(</a><a id="810" href="Relation.Binary.Structures.html#1550" class="Record">IsEquivalence</a><a id="823" class="Symbol">)</a>
<a id="826" class="Keyword">module</a> <a id="833" href="Algebra.Lattice.Structures.html" class="Module">Algebra.Lattice.Structures</a>
<a id="862" class="Symbol">{</a><a id="863" href="Algebra.Lattice.Structures.html#863" class="Bound">a</a> <a id="865" href="Algebra.Lattice.Structures.html#865" class="Bound"></a><a id="866" class="Symbol">}</a> <a id="868" class="Symbol">{</a><a id="869" href="Algebra.Lattice.Structures.html#869" class="Bound">A</a> <a id="871" class="Symbol">:</a> <a id="873" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="877" href="Algebra.Lattice.Structures.html#863" class="Bound">a</a><a id="878" class="Symbol">}</a> <a id="881" class="Comment">-- The underlying set</a>
<a id="905" class="Symbol">(</a><a id="906" href="Algebra.Lattice.Structures.html#906" class="Bound Operator">_≈_</a> <a id="910" class="Symbol">:</a> <a id="912" href="Relation.Binary.Core.html#896" class="Function">Rel</a> <a id="916" href="Algebra.Lattice.Structures.html#869" class="Bound">A</a> <a id="918" href="Algebra.Lattice.Structures.html#865" class="Bound"></a><a id="919" class="Symbol">)</a> <a id="924" class="Comment">-- The underlying equality relation</a>
<a id="962" class="Keyword">where</a>
<a id="969" class="Keyword">open</a> <a id="974" class="Keyword">import</a> <a id="981" href="Algebra.Definitions.html" class="Module">Algebra.Definitions</a> <a id="1001" href="Algebra.Lattice.Structures.html#906" class="Bound Operator">_≈_</a>
<a id="1005" class="Keyword">open</a> <a id="1010" class="Keyword">import</a> <a id="1017" href="Algebra.Structures.html" class="Module">Algebra.Structures</a> <a id="1036" href="Algebra.Lattice.Structures.html#906" class="Bound Operator">_≈_</a>
<a id="1041" class="Comment">------------------------------------------------------------------------</a>
<a id="1114" class="Comment">-- Structures with 1 binary operation</a>
<a id="1153" class="Keyword">record</a> <a id="IsSemilattice"></a><a id="1160" href="Algebra.Lattice.Structures.html#1160" class="Record">IsSemilattice</a> <a id="1174" class="Symbol">(</a><a id="1175" href="Algebra.Lattice.Structures.html#1175" class="Bound"></a> <a id="1177" class="Symbol">:</a> <a id="1179" href="Algebra.Core.html#527" class="Function">Op₂</a> <a id="1183" href="Algebra.Lattice.Structures.html#869" class="Bound">A</a><a id="1184" class="Symbol">)</a> <a id="1186" class="Symbol">:</a> <a id="1188" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1192" class="Symbol">(</a><a id="1193" href="Algebra.Lattice.Structures.html#863" class="Bound">a</a> <a id="1195" href="Agda.Primitive.html#961" class="Primitive Operator"></a> <a id="1197" href="Algebra.Lattice.Structures.html#865" class="Bound"></a><a id="1198" class="Symbol">)</a> <a id="1200" class="Keyword">where</a>
<a id="1208" class="Keyword">field</a>
<a id="IsSemilattice.isBand"></a><a id="1218" href="Algebra.Lattice.Structures.html#1218" class="Field">isBand</a> <a id="1225" class="Symbol">:</a> <a id="1227" href="Algebra.Structures.html#3041" class="Record">IsBand</a> <a id="1234" href="Algebra.Lattice.Structures.html#1175" class="Bound"></a>
<a id="IsSemilattice.comm"></a><a id="1240" href="Algebra.Lattice.Structures.html#1240" class="Field">comm</a> <a id="1247" class="Symbol">:</a> <a id="1249" href="Algebra.Definitions.html#1635" class="Function">Commutative</a> <a id="1261" href="Algebra.Lattice.Structures.html#1175" class="Bound"></a>
<a id="1266" class="Keyword">open</a> <a id="1271" href="Algebra.Structures.html#3041" class="Module">IsBand</a> <a id="1278" href="Algebra.Lattice.Structures.html#1218" class="Field">isBand</a> <a id="1285" class="Keyword">public</a>
<a id="1293" class="Comment">-- Used to bring names appropriate for a meet semilattice into scope.</a>
<a id="IsMeetSemilattice"></a><a id="1363" href="Algebra.Lattice.Structures.html#1363" class="Function">IsMeetSemilattice</a> <a id="1381" class="Symbol">=</a> <a id="1383" href="Algebra.Lattice.Structures.html#1160" class="Record">IsSemilattice</a>
<a id="1397" class="Keyword">module</a> <a id="IsMeetSemilattice"></a><a id="1404" href="Algebra.Lattice.Structures.html#1404" class="Module">IsMeetSemilattice</a> <a id="1422" class="Symbol">{</a><a id="1423" href="Algebra.Lattice.Structures.html#1423" class="Bound"></a><a id="1424" class="Symbol">}</a> <a id="1426" class="Symbol">(</a><a id="1427" href="Algebra.Lattice.Structures.html#1427" class="Bound">L</a> <a id="1429" class="Symbol">:</a> <a id="1431" href="Algebra.Lattice.Structures.html#1363" class="Function">IsMeetSemilattice</a> <a id="1449" href="Algebra.Lattice.Structures.html#1423" class="Bound"></a><a id="1450" class="Symbol">)</a> <a id="1452" class="Keyword">where</a>
<a id="1460" class="Keyword">open</a> <a id="1465" href="Algebra.Lattice.Structures.html#1160" class="Module">IsSemilattice</a> <a id="1479" href="Algebra.Lattice.Structures.html#1427" class="Bound">L</a> <a id="1481" class="Keyword">public</a>
<a id="1492" class="Keyword">renaming</a>
<a id="1505" class="Symbol">(</a> <a id="1507" href="Algebra.Structures.html#1315" class="Function">∙-cong</a> <a id="1515" class="Symbol">to</a> <a id="1518" class="Function">∧-cong</a>
<a id="1529" class="Symbol">;</a> <a id="1531" href="Algebra.Structures.html#1465" class="Function">∙-congˡ</a> <a id="1539" class="Symbol">to</a> <a id="1542" class="Function">∧-congˡ</a>
<a id="1554" class="Symbol">;</a> <a id="1556" href="Algebra.Structures.html#1526" class="Function">∙-congʳ</a> <a id="1564" class="Symbol">to</a> <a id="1567" class="Function">∧-congʳ</a>
<a id="1579" class="Symbol">)</a>
<a id="1582" class="Comment">-- Used to bring names appropriate for a join semilattice into scope.</a>
<a id="IsJoinSemilattice"></a><a id="1652" href="Algebra.Lattice.Structures.html#1652" class="Function">IsJoinSemilattice</a> <a id="1670" class="Symbol">=</a> <a id="1672" href="Algebra.Lattice.Structures.html#1160" class="Record">IsSemilattice</a>
<a id="1686" class="Keyword">module</a> <a id="IsJoinSemilattice"></a><a id="1693" href="Algebra.Lattice.Structures.html#1693" class="Module">IsJoinSemilattice</a> <a id="1711" class="Symbol">{</a><a id="1712" href="Algebra.Lattice.Structures.html#1712" class="Bound"></a><a id="1713" class="Symbol">}</a> <a id="1715" class="Symbol">(</a><a id="1716" href="Algebra.Lattice.Structures.html#1716" class="Bound">L</a> <a id="1718" class="Symbol">:</a> <a id="1720" href="Algebra.Lattice.Structures.html#1652" class="Function">IsJoinSemilattice</a> <a id="1738" href="Algebra.Lattice.Structures.html#1712" class="Bound"></a><a id="1739" class="Symbol">)</a> <a id="1741" class="Keyword">where</a>
<a id="1749" class="Keyword">open</a> <a id="1754" href="Algebra.Lattice.Structures.html#1160" class="Module">IsSemilattice</a> <a id="1768" href="Algebra.Lattice.Structures.html#1716" class="Bound">L</a> <a id="1770" class="Keyword">public</a>
<a id="1781" class="Keyword">renaming</a>
<a id="1794" class="Symbol">(</a> <a id="1796" href="Algebra.Structures.html#1315" class="Function">∙-cong</a> <a id="1804" class="Symbol">to</a> <a id="1807" class="Function">-cong</a>
<a id="1818" class="Symbol">;</a> <a id="1820" href="Algebra.Structures.html#1465" class="Function">∙-congˡ</a> <a id="1828" class="Symbol">to</a> <a id="1831" class="Function">-congˡ</a>
<a id="1843" class="Symbol">;</a> <a id="1845" href="Algebra.Structures.html#1526" class="Function">∙-congʳ</a> <a id="1853" class="Symbol">to</a> <a id="1856" class="Function">-congʳ</a>
<a id="1868" class="Symbol">)</a>
<a id="1871" class="Comment">------------------------------------------------------------------------</a>
<a id="1944" class="Comment">-- Structures with 1 binary operation &amp; 1 element</a>
<a id="1995" class="Comment">-- A bounded semi-lattice is the same thing as an idempotent commutative</a>
<a id="2068" class="Comment">-- monoid.</a>
<a id="IsBoundedSemilattice"></a><a id="2079" href="Algebra.Lattice.Structures.html#2079" class="Function">IsBoundedSemilattice</a> <a id="2100" class="Symbol">=</a> <a id="2102" href="Algebra.Structures.html#4793" class="Record">IsIdempotentCommutativeMonoid</a>
<a id="2132" class="Keyword">module</a> <a id="IsBoundedSemilattice"></a><a id="2139" href="Algebra.Lattice.Structures.html#2139" class="Module">IsBoundedSemilattice</a> <a id="2160" class="Symbol">{</a><a id="2161" href="Algebra.Lattice.Structures.html#2161" class="Bound"></a> <a id="2163" href="Algebra.Lattice.Structures.html#2163" class="Bound">ε</a><a id="2164" class="Symbol">}</a> <a id="2166" class="Symbol">(</a><a id="2167" href="Algebra.Lattice.Structures.html#2167" class="Bound">L</a> <a id="2169" class="Symbol">:</a> <a id="2171" href="Algebra.Lattice.Structures.html#2079" class="Function">IsBoundedSemilattice</a> <a id="2192" href="Algebra.Lattice.Structures.html#2161" class="Bound"></a> <a id="2194" href="Algebra.Lattice.Structures.html#2163" class="Bound">ε</a><a id="2195" class="Symbol">)</a> <a id="2197" class="Keyword">where</a>
<a id="2206" class="Keyword">open</a> <a id="2211" href="Algebra.Structures.html#4793" class="Module">IsIdempotentCommutativeMonoid</a> <a id="2241" href="Algebra.Lattice.Structures.html#2167" class="Bound">L</a> <a id="2243" class="Keyword">public</a>
<a id="IsBoundedSemilattice.isSemilattice"></a><a id="2253" href="Algebra.Lattice.Structures.html#2253" class="Function">isSemilattice</a> <a id="2267" class="Symbol">:</a> <a id="2269" href="Algebra.Lattice.Structures.html#1160" class="Record">IsSemilattice</a> <a id="2283" href="Algebra.Lattice.Structures.html#2161" class="Bound"></a>
<a id="2287" href="Algebra.Lattice.Structures.html#2253" class="Function">isSemilattice</a> <a id="2301" class="Symbol">=</a> <a id="2303" class="Keyword">record</a>
<a id="2314" class="Symbol">{</a> <a id="2316" href="Algebra.Lattice.Structures.html#1218" class="Field">isBand</a> <a id="2323" class="Symbol">=</a> <a id="2325" href="Algebra.Structures.html#5055" class="Function">isBand</a>
<a id="2336" class="Symbol">;</a> <a id="2338" href="Algebra.Lattice.Structures.html#1240" class="Field">comm</a> <a id="2345" class="Symbol">=</a> <a id="2347" href="Algebra.Structures.html#4484" class="Function">comm</a>
<a id="2356" class="Symbol">}</a>
<a id="2360" class="Comment">-- Used to bring names appropriate for a bounded meet semilattice</a>
<a id="2426" class="Comment">-- into scope.</a>
<a id="IsBoundedMeetSemilattice"></a><a id="2441" href="Algebra.Lattice.Structures.html#2441" class="Function">IsBoundedMeetSemilattice</a> <a id="2466" class="Symbol">=</a> <a id="2468" href="Algebra.Lattice.Structures.html#2079" class="Function">IsBoundedSemilattice</a>
<a id="2489" class="Keyword">module</a> <a id="IsBoundedMeetSemilattice"></a><a id="2496" href="Algebra.Lattice.Structures.html#2496" class="Module">IsBoundedMeetSemilattice</a> <a id="2521" class="Symbol">{</a><a id="2522" href="Algebra.Lattice.Structures.html#2522" class="Bound"></a> <a id="2524" href="Algebra.Lattice.Structures.html#2524" class="Bound"></a><a id="2525" class="Symbol">}</a> <a id="2527" class="Symbol">(</a><a id="2528" href="Algebra.Lattice.Structures.html#2528" class="Bound">L</a> <a id="2530" class="Symbol">:</a> <a id="2532" href="Algebra.Lattice.Structures.html#2441" class="Function">IsBoundedMeetSemilattice</a> <a id="2557" href="Algebra.Lattice.Structures.html#2522" class="Bound"></a> <a id="2559" href="Algebra.Lattice.Structures.html#2524" class="Bound"></a><a id="2560" class="Symbol">)</a>
<a id="2564" class="Keyword">where</a>
<a id="2573" class="Keyword">open</a> <a id="2578" href="Algebra.Lattice.Structures.html#2139" class="Module">IsBoundedSemilattice</a> <a id="2599" href="Algebra.Lattice.Structures.html#2528" class="Bound">L</a> <a id="2601" class="Keyword">public</a>
<a id="2612" class="Keyword">using</a> <a id="2618" class="Symbol">(</a><a id="2619" href="Algebra.Structures.html#4067" class="Function">identity</a><a id="2627" class="Symbol">;</a> <a id="2629" href="Algebra.Structures.html#4136" class="Function">identityˡ</a><a id="2638" class="Symbol">;</a> <a id="2640" href="Algebra.Structures.html#4197" class="Function">identityʳ</a><a id="2649" class="Symbol">)</a>
<a id="2655" class="Keyword">renaming</a> <a id="2664" class="Symbol">(</a><a id="2665" href="Algebra.Lattice.Structures.html#2253" class="Function">isSemilattice</a> <a id="2679" class="Symbol">to</a> <a id="2682" class="Function">isMeetSemilattice</a><a id="2699" class="Symbol">)</a>
<a id="2704" class="Keyword">open</a> <a id="2709" href="Algebra.Lattice.Structures.html#1404" class="Module">IsMeetSemilattice</a> <a id="2727" href="Algebra.Lattice.Structures.html#2682" class="Function">isMeetSemilattice</a> <a id="2745" class="Keyword">public</a>
<a id="2754" class="Comment">-- Used to bring names appropriate for a bounded join semilattice</a>
<a id="2820" class="Comment">-- into scope.</a>
<a id="IsBoundedJoinSemilattice"></a><a id="2835" href="Algebra.Lattice.Structures.html#2835" class="Function">IsBoundedJoinSemilattice</a> <a id="2860" class="Symbol">=</a> <a id="2862" href="Algebra.Lattice.Structures.html#2079" class="Function">IsBoundedSemilattice</a>
<a id="2883" class="Keyword">module</a> <a id="IsBoundedJoinSemilattice"></a><a id="2890" href="Algebra.Lattice.Structures.html#2890" class="Module">IsBoundedJoinSemilattice</a> <a id="2915" class="Symbol">{</a><a id="2916" href="Algebra.Lattice.Structures.html#2916" class="Bound"></a> <a id="2918" href="Algebra.Lattice.Structures.html#2918" class="Bound"></a><a id="2919" class="Symbol">}</a> <a id="2921" class="Symbol">(</a><a id="2922" href="Algebra.Lattice.Structures.html#2922" class="Bound">L</a> <a id="2924" class="Symbol">:</a> <a id="2926" href="Algebra.Lattice.Structures.html#2835" class="Function">IsBoundedJoinSemilattice</a> <a id="2951" href="Algebra.Lattice.Structures.html#2916" class="Bound"></a> <a id="2953" href="Algebra.Lattice.Structures.html#2918" class="Bound"></a><a id="2954" class="Symbol">)</a>
<a id="2958" class="Keyword">where</a>
<a id="2967" class="Keyword">open</a> <a id="2972" href="Algebra.Lattice.Structures.html#2139" class="Module">IsBoundedSemilattice</a> <a id="2993" href="Algebra.Lattice.Structures.html#2922" class="Bound">L</a> <a id="2995" class="Keyword">public</a>
<a id="3006" class="Keyword">using</a> <a id="3012" class="Symbol">(</a><a id="3013" href="Algebra.Structures.html#4067" class="Function">identity</a><a id="3021" class="Symbol">;</a> <a id="3023" href="Algebra.Structures.html#4136" class="Function">identityˡ</a><a id="3032" class="Symbol">;</a> <a id="3034" href="Algebra.Structures.html#4197" class="Function">identityʳ</a><a id="3043" class="Symbol">)</a>
<a id="3049" class="Keyword">renaming</a> <a id="3058" class="Symbol">(</a><a id="3059" href="Algebra.Lattice.Structures.html#2253" class="Function">isSemilattice</a> <a id="3073" class="Symbol">to</a> <a id="3076" class="Function">isJoinSemilattice</a><a id="3093" class="Symbol">)</a>
<a id="3098" class="Keyword">open</a> <a id="3103" href="Algebra.Lattice.Structures.html#1693" class="Module">IsJoinSemilattice</a> <a id="3121" href="Algebra.Lattice.Structures.html#3076" class="Function">isJoinSemilattice</a> <a id="3139" class="Keyword">public</a>
<a id="3147" class="Comment">------------------------------------------------------------------------</a>
<a id="3220" class="Comment">-- Structures with 2 binary operations</a>
<a id="3260" class="Comment">-- Note that `IsLattice` is not defined in terms of `IsMeetSemilattice`</a>
<a id="3332" class="Comment">-- and `IsJoinSemilattice` for two reasons:</a>
<a id="3376" class="Comment">-- 1) it would result in a structure with two *different* proofs that</a>
<a id="3448" class="Comment">-- the equality relation `≈` is an equivalence relation.</a>
<a id="3507" class="Comment">-- 2) the idempotence laws of and ∧ can be derived from the</a>
<a id="3571" class="Comment">-- absorption laws, which makes the corresponding &quot;idem&quot; fields</a>
<a id="3637" class="Comment">-- redundant.</a>
<a id="3653" class="Comment">--</a>
<a id="3656" class="Comment">-- It is possible to construct the `IsLattice` record from</a>
<a id="3715" class="Comment">-- `IsMeetSemilattice` and `IsJoinSemilattice` via the `IsLattice₂`</a>
<a id="3783" class="Comment">-- record found in `Algebra.Lattice.Structures.Biased`.</a>
<a id="3839" class="Comment">--</a>
<a id="3842" class="Comment">-- The derived idempotence laws are stated and proved in</a>
<a id="3899" class="Comment">-- `Algebra.Lattice.Properties.Lattice` along with the fact that every</a>
<a id="3970" class="Comment">-- lattice consists of two semilattices.</a>
<a id="4012" class="Keyword">record</a> <a id="IsLattice"></a><a id="4019" href="Algebra.Lattice.Structures.html#4019" class="Record">IsLattice</a> <a id="4029" class="Symbol">(</a><a id="4030" href="Algebra.Lattice.Structures.html#4030" class="Bound"></a> <a id="4032" href="Algebra.Lattice.Structures.html#4032" class="Bound"></a> <a id="4034" class="Symbol">:</a> <a id="4036" href="Algebra.Core.html#527" class="Function">Op₂</a> <a id="4040" href="Algebra.Lattice.Structures.html#869" class="Bound">A</a><a id="4041" class="Symbol">)</a> <a id="4043" class="Symbol">:</a> <a id="4045" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="4049" class="Symbol">(</a><a id="4050" href="Algebra.Lattice.Structures.html#863" class="Bound">a</a> <a id="4052" href="Agda.Primitive.html#961" class="Primitive Operator"></a> <a id="4054" href="Algebra.Lattice.Structures.html#865" class="Bound"></a><a id="4055" class="Symbol">)</a> <a id="4057" class="Keyword">where</a>
<a id="4065" class="Keyword">field</a>
<a id="IsLattice.isEquivalence"></a><a id="4075" href="Algebra.Lattice.Structures.html#4075" class="Field">isEquivalence</a> <a id="4089" class="Symbol">:</a> <a id="4091" href="Relation.Binary.Structures.html#1550" class="Record">IsEquivalence</a> <a id="4105" href="Algebra.Lattice.Structures.html#906" class="Bound Operator">_≈_</a>
<a id="IsLattice.-comm"></a><a id="4113" href="Algebra.Lattice.Structures.html#4113" class="Field">-comm</a> <a id="4127" class="Symbol">:</a> <a id="4129" href="Algebra.Definitions.html#1635" class="Function">Commutative</a> <a id="4141" href="Algebra.Lattice.Structures.html#4030" class="Bound"></a>
<a id="IsLattice.-assoc"></a><a id="4147" href="Algebra.Lattice.Structures.html#4147" class="Field">-assoc</a> <a id="4161" class="Symbol">:</a> <a id="4163" href="Algebra.Definitions.html#1548" class="Function">Associative</a> <a id="4175" href="Algebra.Lattice.Structures.html#4030" class="Bound"></a>
<a id="IsLattice.-cong"></a><a id="4181" href="Algebra.Lattice.Structures.html#4181" class="Field">-cong</a> <a id="4195" class="Symbol">:</a> <a id="4197" href="Algebra.Definitions.html#1302" class="Function">Congruent₂</a> <a id="4208" href="Algebra.Lattice.Structures.html#4030" class="Bound"></a>
<a id="IsLattice.∧-comm"></a><a id="4214" href="Algebra.Lattice.Structures.html#4214" class="Field">∧-comm</a> <a id="4228" class="Symbol">:</a> <a id="4230" href="Algebra.Definitions.html#1635" class="Function">Commutative</a> <a id="4242" href="Algebra.Lattice.Structures.html#4032" class="Bound"></a>
<a id="IsLattice.∧-assoc"></a><a id="4248" href="Algebra.Lattice.Structures.html#4248" class="Field">∧-assoc</a> <a id="4262" class="Symbol">:</a> <a id="4264" href="Algebra.Definitions.html#1548" class="Function">Associative</a> <a id="4276" href="Algebra.Lattice.Structures.html#4032" class="Bound"></a>
<a id="IsLattice.∧-cong"></a><a id="4282" href="Algebra.Lattice.Structures.html#4282" class="Field">∧-cong</a> <a id="4296" class="Symbol">:</a> <a id="4298" href="Algebra.Definitions.html#1302" class="Function">Congruent₂</a> <a id="4309" href="Algebra.Lattice.Structures.html#4032" class="Bound"></a>
<a id="IsLattice.absorptive"></a><a id="4315" href="Algebra.Lattice.Structures.html#4315" class="Field">absorptive</a> <a id="4329" class="Symbol">:</a> <a id="4331" href="Algebra.Definitions.html#3997" class="Function">Absorptive</a> <a id="4342" href="Algebra.Lattice.Structures.html#4030" class="Bound"></a> <a id="4344" href="Algebra.Lattice.Structures.html#4032" class="Bound"></a>
<a id="4349" class="Keyword">open</a> <a id="4354" href="Relation.Binary.Structures.html#1550" class="Module">IsEquivalence</a> <a id="4368" href="Algebra.Lattice.Structures.html#4075" class="Field">isEquivalence</a> <a id="4382" class="Keyword">public</a>
<a id="IsLattice.-absorbs-∧"></a><a id="4392" href="Algebra.Lattice.Structures.html#4392" class="Function">-absorbs-∧</a> <a id="4404" class="Symbol">:</a> <a id="4406" href="Algebra.Lattice.Structures.html#4030" class="Bound"></a> <a id="4408" href="Algebra.Definitions.html#3918" class="Function Operator">Absorbs</a> <a id="4416" href="Algebra.Lattice.Structures.html#4032" class="Bound"></a>
<a id="4420" href="Algebra.Lattice.Structures.html#4392" class="Function">-absorbs-∧</a> <a id="4432" class="Symbol">=</a> <a id="4434" href="Data.Product.Base.html#636" class="Field">proj₁</a> <a id="4440" href="Algebra.Lattice.Structures.html#4315" class="Field">absorptive</a>
<a id="IsLattice.∧-absorbs-"></a><a id="4454" href="Algebra.Lattice.Structures.html#4454" class="Function">∧-absorbs-</a> <a id="4466" class="Symbol">:</a> <a id="4468" href="Algebra.Lattice.Structures.html#4032" class="Bound"></a> <a id="4470" href="Algebra.Definitions.html#3918" class="Function Operator">Absorbs</a> <a id="4478" href="Algebra.Lattice.Structures.html#4030" class="Bound"></a>
<a id="4482" href="Algebra.Lattice.Structures.html#4454" class="Function">∧-absorbs-</a> <a id="4494" class="Symbol">=</a> <a id="4496" href="Data.Product.Base.html#650" class="Field">proj₂</a> <a id="4502" href="Algebra.Lattice.Structures.html#4315" class="Field">absorptive</a>
<a id="IsLattice.∧-congˡ"></a><a id="4516" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="4524" class="Symbol">:</a> <a id="4526" href="Algebra.Definitions.html#1374" class="Function">LeftCongruent</a> <a id="4540" href="Algebra.Lattice.Structures.html#4032" class="Bound"></a>
<a id="4544" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="4552" href="Algebra.Lattice.Structures.html#4552" class="Bound">y≈z</a> <a id="4556" class="Symbol">=</a> <a id="4558" href="Algebra.Lattice.Structures.html#4282" class="Field">∧-cong</a> <a id="4565" href="Relation.Binary.Structures.html#1596" class="Function">refl</a> <a id="4570" href="Algebra.Lattice.Structures.html#4552" class="Bound">y≈z</a>
<a id="IsLattice.∧-congʳ"></a><a id="4577" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="4585" class="Symbol">:</a> <a id="4587" href="Algebra.Definitions.html#1460" class="Function">RightCongruent</a> <a id="4602" href="Algebra.Lattice.Structures.html#4032" class="Bound"></a>
<a id="4606" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="4614" href="Algebra.Lattice.Structures.html#4614" class="Bound">y≈z</a> <a id="4618" class="Symbol">=</a> <a id="4620" href="Algebra.Lattice.Structures.html#4282" class="Field">∧-cong</a> <a id="4627" href="Algebra.Lattice.Structures.html#4614" class="Bound">y≈z</a> <a id="4631" href="Relation.Binary.Structures.html#1596" class="Function">refl</a>
<a id="IsLattice.-congˡ"></a><a id="4639" href="Algebra.Lattice.Structures.html#4639" class="Function">-congˡ</a> <a id="4647" class="Symbol">:</a> <a id="4649" href="Algebra.Definitions.html#1374" class="Function">LeftCongruent</a> <a id="4663" href="Algebra.Lattice.Structures.html#4030" class="Bound"></a>
<a id="4667" href="Algebra.Lattice.Structures.html#4639" class="Function">-congˡ</a> <a id="4675" href="Algebra.Lattice.Structures.html#4675" class="Bound">y≈z</a> <a id="4679" class="Symbol">=</a> <a id="4681" href="Algebra.Lattice.Structures.html#4181" class="Field">-cong</a> <a id="4688" href="Relation.Binary.Structures.html#1596" class="Function">refl</a> <a id="4693" href="Algebra.Lattice.Structures.html#4675" class="Bound">y≈z</a>
<a id="IsLattice.-congʳ"></a><a id="4700" href="Algebra.Lattice.Structures.html#4700" class="Function">-congʳ</a> <a id="4708" class="Symbol">:</a> <a id="4710" href="Algebra.Definitions.html#1460" class="Function">RightCongruent</a> <a id="4725" href="Algebra.Lattice.Structures.html#4030" class="Bound"></a>
<a id="4729" href="Algebra.Lattice.Structures.html#4700" class="Function">-congʳ</a> <a id="4737" href="Algebra.Lattice.Structures.html#4737" class="Bound">y≈z</a> <a id="4741" class="Symbol">=</a> <a id="4743" href="Algebra.Lattice.Structures.html#4181" class="Field">-cong</a> <a id="4750" href="Algebra.Lattice.Structures.html#4737" class="Bound">y≈z</a> <a id="4754" href="Relation.Binary.Structures.html#1596" class="Function">refl</a>
<a id="4761" class="Keyword">record</a> <a id="IsDistributiveLattice"></a><a id="4768" href="Algebra.Lattice.Structures.html#4768" class="Record">IsDistributiveLattice</a> <a id="4790" class="Symbol">(</a><a id="4791" href="Algebra.Lattice.Structures.html#4791" class="Bound"></a> <a id="4793" href="Algebra.Lattice.Structures.html#4793" class="Bound"></a> <a id="4795" class="Symbol">:</a> <a id="4797" href="Algebra.Core.html#527" class="Function">Op₂</a> <a id="4801" href="Algebra.Lattice.Structures.html#869" class="Bound">A</a><a id="4802" class="Symbol">)</a> <a id="4804" class="Symbol">:</a> <a id="4806" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="4810" class="Symbol">(</a><a id="4811" href="Algebra.Lattice.Structures.html#863" class="Bound">a</a> <a id="4813" href="Agda.Primitive.html#961" class="Primitive Operator"></a> <a id="4815" href="Algebra.Lattice.Structures.html#865" class="Bound"></a><a id="4816" class="Symbol">)</a> <a id="4818" class="Keyword">where</a>
<a id="4826" class="Keyword">field</a>
<a id="IsDistributiveLattice.isLattice"></a><a id="4836" href="Algebra.Lattice.Structures.html#4836" class="Field">isLattice</a> <a id="4848" class="Symbol">:</a> <a id="4850" href="Algebra.Lattice.Structures.html#4019" class="Record">IsLattice</a> <a id="4860" href="Algebra.Lattice.Structures.html#4791" class="Bound"></a> <a id="4862" href="Algebra.Lattice.Structures.html#4793" class="Bound"></a>
<a id="IsDistributiveLattice.-distrib-∧"></a><a id="4868" href="Algebra.Lattice.Structures.html#4868" class="Field">-distrib-∧</a> <a id="4880" class="Symbol">:</a> <a id="4882" href="Algebra.Lattice.Structures.html#4791" class="Bound"></a> <a id="4884" href="Algebra.Definitions.html#3393" class="Function Operator">DistributesOver</a> <a id="4900" href="Algebra.Lattice.Structures.html#4793" class="Bound"></a>
<a id="IsDistributiveLattice.∧-distrib-"></a><a id="4906" href="Algebra.Lattice.Structures.html#4906" class="Field">∧-distrib-</a> <a id="4918" class="Symbol">:</a> <a id="4920" href="Algebra.Lattice.Structures.html#4793" class="Bound"></a> <a id="4922" href="Algebra.Definitions.html#3393" class="Function Operator">DistributesOver</a> <a id="4938" href="Algebra.Lattice.Structures.html#4791" class="Bound"></a>
<a id="4943" class="Keyword">open</a> <a id="4948" href="Algebra.Lattice.Structures.html#4019" class="Module">IsLattice</a> <a id="4958" href="Algebra.Lattice.Structures.html#4836" class="Field">isLattice</a> <a id="4968" class="Keyword">public</a>
<a id="IsDistributiveLattice.-distribˡ-∧"></a><a id="4978" href="Algebra.Lattice.Structures.html#4978" class="Function">-distribˡ-∧</a> <a id="4991" class="Symbol">:</a> <a id="4993" href="Algebra.Lattice.Structures.html#4791" class="Bound"></a> <a id="4995" href="Algebra.Definitions.html#3155" class="Function Operator">DistributesOverˡ</a> <a id="5012" href="Algebra.Lattice.Structures.html#4793" class="Bound"></a>
<a id="5016" href="Algebra.Lattice.Structures.html#4978" class="Function">-distribˡ-∧</a> <a id="5029" class="Symbol">=</a> <a id="5031" href="Data.Product.Base.html#636" class="Field">proj₁</a> <a id="5037" href="Algebra.Lattice.Structures.html#4868" class="Field">-distrib-∧</a>
<a id="IsDistributiveLattice.-distribʳ-∧"></a><a id="5052" href="Algebra.Lattice.Structures.html#5052" class="Function">-distribʳ-∧</a> <a id="5065" class="Symbol">:</a> <a id="5067" href="Algebra.Lattice.Structures.html#4791" class="Bound"></a> <a id="5069" href="Algebra.Definitions.html#3274" class="Function Operator">DistributesOverʳ</a> <a id="5086" href="Algebra.Lattice.Structures.html#4793" class="Bound"></a>
<a id="5090" href="Algebra.Lattice.Structures.html#5052" class="Function">-distribʳ-∧</a> <a id="5103" class="Symbol">=</a> <a id="5105" href="Data.Product.Base.html#650" class="Field">proj₂</a> <a id="5111" href="Algebra.Lattice.Structures.html#4868" class="Field">-distrib-∧</a>
<a id="IsDistributiveLattice.∧-distribˡ-"></a><a id="5126" href="Algebra.Lattice.Structures.html#5126" class="Function">∧-distribˡ-</a> <a id="5139" class="Symbol">:</a> <a id="5141" href="Algebra.Lattice.Structures.html#4793" class="Bound"></a> <a id="5143" href="Algebra.Definitions.html#3155" class="Function Operator">DistributesOverˡ</a> <a id="5160" href="Algebra.Lattice.Structures.html#4791" class="Bound"></a>
<a id="5164" href="Algebra.Lattice.Structures.html#5126" class="Function">∧-distribˡ-</a> <a id="5177" class="Symbol">=</a> <a id="5179" href="Data.Product.Base.html#636" class="Field">proj₁</a> <a id="5185" href="Algebra.Lattice.Structures.html#4906" class="Field">∧-distrib-</a>
<a id="IsDistributiveLattice.∧-distribʳ-"></a><a id="5200" href="Algebra.Lattice.Structures.html#5200" class="Function">∧-distribʳ-</a> <a id="5213" class="Symbol">:</a> <a id="5215" href="Algebra.Lattice.Structures.html#4793" class="Bound"></a> <a id="5217" href="Algebra.Definitions.html#3274" class="Function Operator">DistributesOverʳ</a> <a id="5234" href="Algebra.Lattice.Structures.html#4791" class="Bound"></a>
<a id="5238" href="Algebra.Lattice.Structures.html#5200" class="Function">∧-distribʳ-</a> <a id="5251" class="Symbol">=</a> <a id="5253" href="Data.Product.Base.html#650" class="Field">proj₂</a> <a id="5259" href="Algebra.Lattice.Structures.html#4906" class="Field">∧-distrib-</a>
<a id="5272" class="Comment">------------------------------------------------------------------------</a>
<a id="5345" class="Comment">-- Structures with 2 binary ops, 1 unary op and 2 elements.</a>
<a id="5406" class="Keyword">record</a> <a id="IsBooleanAlgebra"></a><a id="5413" href="Algebra.Lattice.Structures.html#5413" class="Record">IsBooleanAlgebra</a> <a id="5430" class="Symbol">(</a><a id="5431" href="Algebra.Lattice.Structures.html#5431" class="Bound"></a> <a id="5433" href="Algebra.Lattice.Structures.html#5433" class="Bound"></a> <a id="5435" class="Symbol">:</a> <a id="5437" href="Algebra.Core.html#527" class="Function">Op₂</a> <a id="5441" href="Algebra.Lattice.Structures.html#869" class="Bound">A</a><a id="5442" class="Symbol">)</a> <a id="5444" class="Symbol">(</a><a id="5445" href="Algebra.Lattice.Structures.html#5445" class="Bound">¬</a> <a id="5447" class="Symbol">:</a> <a id="5449" href="Algebra.Core.html#484" class="Function">Op₁</a> <a id="5453" href="Algebra.Lattice.Structures.html#869" class="Bound">A</a><a id="5454" class="Symbol">)</a> <a id="5456" class="Symbol">(</a><a id="5457" href="Algebra.Lattice.Structures.html#5457" class="Bound"></a> <a id="5459" href="Algebra.Lattice.Structures.html#5459" class="Bound"></a> <a id="5461" class="Symbol">:</a> <a id="5463" href="Algebra.Lattice.Structures.html#869" class="Bound">A</a><a id="5464" class="Symbol">)</a> <a id="5466" class="Symbol">:</a> <a id="5468" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="5472" class="Symbol">(</a><a id="5473" href="Algebra.Lattice.Structures.html#863" class="Bound">a</a> <a id="5475" href="Agda.Primitive.html#961" class="Primitive Operator"></a> <a id="5477" href="Algebra.Lattice.Structures.html#865" class="Bound"></a><a id="5478" class="Symbol">)</a>
<a id="5482" class="Keyword">where</a>
<a id="5491" class="Keyword">field</a>
<a id="IsBooleanAlgebra.isDistributiveLattice"></a><a id="5501" href="Algebra.Lattice.Structures.html#5501" class="Field">isDistributiveLattice</a> <a id="5523" class="Symbol">:</a> <a id="5525" href="Algebra.Lattice.Structures.html#4768" class="Record">IsDistributiveLattice</a> <a id="5547" href="Algebra.Lattice.Structures.html#5431" class="Bound"></a> <a id="5549" href="Algebra.Lattice.Structures.html#5433" class="Bound"></a>
<a id="IsBooleanAlgebra.-complement"></a><a id="5555" href="Algebra.Lattice.Structures.html#5555" class="Field">-complement</a> <a id="5577" class="Symbol">:</a> <a id="5579" href="Algebra.Definitions.html#2322" class="Function">Inverse</a> <a id="5587" href="Algebra.Lattice.Structures.html#5457" class="Bound"></a> <a id="5589" href="Algebra.Lattice.Structures.html#5445" class="Bound">¬</a> <a id="5591" href="Algebra.Lattice.Structures.html#5431" class="Bound"></a>
<a id="IsBooleanAlgebra.∧-complement"></a><a id="5597" href="Algebra.Lattice.Structures.html#5597" class="Field">∧-complement</a> <a id="5619" class="Symbol">:</a> <a id="5621" href="Algebra.Definitions.html#2322" class="Function">Inverse</a> <a id="5629" href="Algebra.Lattice.Structures.html#5459" class="Bound"></a> <a id="5631" href="Algebra.Lattice.Structures.html#5445" class="Bound">¬</a> <a id="5633" href="Algebra.Lattice.Structures.html#5433" class="Bound"></a>
<a id="IsBooleanAlgebra.¬-cong"></a><a id="5639" href="Algebra.Lattice.Structures.html#5639" class="Field">¬-cong</a> <a id="5661" class="Symbol">:</a> <a id="5663" href="Algebra.Definitions.html#1237" class="Function">Congruent₁</a> <a id="5674" href="Algebra.Lattice.Structures.html#5445" class="Bound">¬</a>
<a id="5679" class="Keyword">open</a> <a id="5684" href="Algebra.Lattice.Structures.html#4768" class="Module">IsDistributiveLattice</a> <a id="5706" href="Algebra.Lattice.Structures.html#5501" class="Field">isDistributiveLattice</a> <a id="5728" class="Keyword">public</a>
<a id="IsBooleanAlgebra.-complementˡ"></a><a id="5738" href="Algebra.Lattice.Structures.html#5738" class="Function">-complementˡ</a> <a id="5752" class="Symbol">:</a> <a id="5754" href="Algebra.Definitions.html#2144" class="Function">LeftInverse</a> <a id="5766" href="Algebra.Lattice.Structures.html#5457" class="Bound"></a> <a id="5768" href="Algebra.Lattice.Structures.html#5445" class="Bound">¬</a> <a id="5770" href="Algebra.Lattice.Structures.html#5431" class="Bound"></a>
<a id="5774" href="Algebra.Lattice.Structures.html#5738" class="Function">-complementˡ</a> <a id="5788" class="Symbol">=</a> <a id="5790" href="Data.Product.Base.html#636" class="Field">proj₁</a> <a id="5796" href="Algebra.Lattice.Structures.html#5555" class="Field">-complement</a>
<a id="IsBooleanAlgebra.-complementʳ"></a><a id="5812" href="Algebra.Lattice.Structures.html#5812" class="Function">-complementʳ</a> <a id="5826" class="Symbol">:</a> <a id="5828" href="Algebra.Definitions.html#2232" class="Function">RightInverse</a> <a id="5841" href="Algebra.Lattice.Structures.html#5457" class="Bound"></a> <a id="5843" href="Algebra.Lattice.Structures.html#5445" class="Bound">¬</a> <a id="5845" href="Algebra.Lattice.Structures.html#5431" class="Bound"></a>
<a id="5849" href="Algebra.Lattice.Structures.html#5812" class="Function">-complementʳ</a> <a id="5863" class="Symbol">=</a> <a id="5865" href="Data.Product.Base.html#650" class="Field">proj₂</a> <a id="5871" href="Algebra.Lattice.Structures.html#5555" class="Field">-complement</a>
<a id="IsBooleanAlgebra.∧-complementˡ"></a><a id="5887" href="Algebra.Lattice.Structures.html#5887" class="Function">∧-complementˡ</a> <a id="5901" class="Symbol">:</a> <a id="5903" href="Algebra.Definitions.html#2144" class="Function">LeftInverse</a> <a id="5915" href="Algebra.Lattice.Structures.html#5459" class="Bound"></a> <a id="5917" href="Algebra.Lattice.Structures.html#5445" class="Bound">¬</a> <a id="5919" href="Algebra.Lattice.Structures.html#5433" class="Bound"></a>
<a id="5923" href="Algebra.Lattice.Structures.html#5887" class="Function">∧-complementˡ</a> <a id="5937" class="Symbol">=</a> <a id="5939" href="Data.Product.Base.html#636" class="Field">proj₁</a> <a id="5945" href="Algebra.Lattice.Structures.html#5597" class="Field">∧-complement</a>
<a id="IsBooleanAlgebra.∧-complementʳ"></a><a id="5961" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="5975" class="Symbol">:</a> <a id="5977" href="Algebra.Definitions.html#2232" class="Function">RightInverse</a> <a id="5990" href="Algebra.Lattice.Structures.html#5459" class="Bound"></a> <a id="5992" href="Algebra.Lattice.Structures.html#5445" class="Bound">¬</a> <a id="5994" href="Algebra.Lattice.Structures.html#5433" class="Bound"></a>
<a id="5998" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="6012" class="Symbol">=</a> <a id="6014" href="Data.Product.Base.html#650" class="Field">proj₂</a> <a id="6020" href="Algebra.Lattice.Structures.html#5597" class="Field">∧-complement</a>
</pre></body></html>