mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
147 lines
56 KiB
HTML
147 lines
56 KiB
HTML
|
<!DOCTYPE HTML>
|
||
|
<html><head><meta charset="utf-8"><title>Algebra.Construct.NaturalChoice.MinMaxOp</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">-- Properties of min and max operators specified over a total</a>
|
||
|
<a id="168" class="Comment">-- preorder.</a>
|
||
|
<a id="181" class="Comment">------------------------------------------------------------------------</a>
|
||
|
|
||
|
<a id="255" class="Symbol">{-#</a> <a id="259" class="Keyword">OPTIONS</a> <a id="267" class="Pragma">--cubical-compatible</a> <a id="288" class="Pragma">--safe</a> <a id="295" class="Symbol">#-}</a>
|
||
|
|
||
|
<a id="300" class="Keyword">open</a> <a id="305" class="Keyword">import</a> <a id="312" href="Algebra.Core.html" class="Module">Algebra.Core</a>
|
||
|
<a id="325" class="Keyword">open</a> <a id="330" class="Keyword">import</a> <a id="337" href="Algebra.Bundles.html" class="Module">Algebra.Bundles</a>
|
||
|
<a id="353" class="Keyword">open</a> <a id="358" class="Keyword">import</a> <a id="365" href="Algebra.Construct.NaturalChoice.Base.html" class="Module">Algebra.Construct.NaturalChoice.Base</a>
|
||
|
<a id="402" class="Keyword">open</a> <a id="407" class="Keyword">import</a> <a id="414" href="Data.Sum.Base.html" class="Module">Data.Sum.Base</a> <a id="428" class="Symbol">as</a> <a id="431" class="Module">Sum</a> <a id="435" class="Keyword">using</a> <a id="441" class="Symbol">(</a><a id="442" href="Data.Sum.Base.html#675" class="InductiveConstructor">inj₁</a><a id="446" class="Symbol">;</a> <a id="448" href="Data.Sum.Base.html#700" class="InductiveConstructor">inj₂</a><a id="452" class="Symbol">;</a> <a id="454" href="Data.Sum.Base.html#811" class="Function Operator">[_,_]</a><a id="459" class="Symbol">)</a>
|
||
|
<a id="461" class="Keyword">open</a> <a id="466" class="Keyword">import</a> <a id="473" href="Data.Product.Base.html" class="Module">Data.Product.Base</a> <a id="491" class="Keyword">using</a> <a id="497" class="Symbol">(</a><a id="498" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">_,_</a><a id="501" class="Symbol">)</a>
|
||
|
<a id="503" class="Keyword">open</a> <a id="508" class="Keyword">import</a> <a id="515" href="Function.Base.html" class="Module">Function.Base</a> <a id="529" class="Keyword">using</a> <a id="535" class="Symbol">(</a><a id="536" href="Function.Base.html#704" class="Function">id</a><a id="538" class="Symbol">;</a> <a id="540" href="Function.Base.html#1115" class="Function Operator">_∘_</a><a id="543" class="Symbol">;</a> <a id="545" href="Function.Base.html#1638" class="Function">flip</a><a id="549" class="Symbol">)</a>
|
||
|
<a id="551" class="Keyword">open</a> <a id="556" class="Keyword">import</a> <a id="563" href="Relation.Binary.Core.html" class="Module">Relation.Binary.Core</a> <a id="584" class="Keyword">using</a> <a id="590" class="Symbol">(</a><a id="591" href="Relation.Binary.Core.html#1577" class="Function Operator">_Preserves_⟶_</a><a id="604" class="Symbol">)</a>
|
||
|
<a id="606" class="Keyword">open</a> <a id="611" class="Keyword">import</a> <a id="618" href="Relation.Binary.Bundles.html" class="Module">Relation.Binary.Bundles</a> <a id="642" class="Keyword">using</a> <a id="648" class="Symbol">(</a><a id="649" href="Relation.Binary.Bundles.html#2842" class="Record">TotalPreorder</a><a id="662" class="Symbol">)</a>
|
||
|
<a id="664" class="Keyword">open</a> <a id="669" class="Keyword">import</a> <a id="676" href="Relation.Binary.Consequences.html" class="Module">Relation.Binary.Consequences</a>
|
||
|
|
||
|
<a id="706" class="Keyword">module</a> <a id="713" href="Algebra.Construct.NaturalChoice.MinMaxOp.html" class="Module">Algebra.Construct.NaturalChoice.MinMaxOp</a>
|
||
|
<a id="756" class="Symbol">{</a><a id="757" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#757" class="Bound">a</a> <a id="759" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#759" class="Bound">ℓ₁</a> <a id="762" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#762" class="Bound">ℓ₂</a><a id="764" class="Symbol">}</a> <a id="766" class="Symbol">{</a><a id="767" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#767" class="Bound">O</a> <a id="769" class="Symbol">:</a> <a id="771" href="Relation.Binary.Bundles.html#2842" class="Record">TotalPreorder</a> <a id="785" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#757" class="Bound">a</a> <a id="787" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#759" class="Bound">ℓ₁</a> <a id="790" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#762" class="Bound">ℓ₂</a><a id="792" class="Symbol">}</a>
|
||
|
<a id="796" class="Symbol">(</a><a id="797" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#797" class="Bound">minOp</a> <a id="803" class="Symbol">:</a> <a id="805" href="Algebra.Construct.NaturalChoice.Base.html#990" class="Record">MinOperator</a> <a id="817" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#767" class="Bound">O</a><a id="818" class="Symbol">)</a>
|
||
|
<a id="822" class="Symbol">(</a><a id="823" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#823" class="Bound">maxOp</a> <a id="829" class="Symbol">:</a> <a id="831" href="Algebra.Construct.NaturalChoice.Base.html#1191" class="Record">MaxOperator</a> <a id="843" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#767" class="Bound">O</a><a id="844" class="Symbol">)</a>
|
||
|
<a id="848" class="Keyword">where</a>
|
||
|
|
||
|
<a id="855" class="Keyword">open</a> <a id="860" href="Relation.Binary.Bundles.html#2842" class="Module">TotalPreorder</a> <a id="874" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#767" class="Bound">O</a> <a id="876" class="Keyword">renaming</a>
|
||
|
<a id="887" class="Symbol">(</a> <a id="889" href="Relation.Binary.Bundles.html#2926" class="Field">Carrier</a> <a id="899" class="Symbol">to</a> <a id="902" class="Field">A</a>
|
||
|
<a id="906" class="Symbol">;</a> <a id="908" href="Relation.Binary.Bundles.html#3020" class="Field Operator">_≲_</a> <a id="918" class="Symbol">to</a> <a id="921" class="Field Operator">_≤_</a>
|
||
|
<a id="927" class="Symbol">;</a> <a id="929" href="Relation.Binary.Structures.html#2688" class="Function">≲-resp-≈</a> <a id="939" class="Symbol">to</a> <a id="942" class="Function">≤-resp-≈</a>
|
||
|
<a id="953" class="Symbol">;</a> <a id="955" href="Relation.Binary.Structures.html#2607" class="Function">≲-respʳ-≈</a> <a id="965" class="Symbol">to</a> <a id="968" class="Function">≤-respʳ-≈</a>
|
||
|
<a id="980" class="Symbol">;</a> <a id="982" href="Relation.Binary.Structures.html#2517" class="Function">≲-respˡ-≈</a> <a id="992" class="Symbol">to</a> <a id="995" class="Function">≤-respˡ-≈</a>
|
||
|
<a id="1007" class="Symbol">)</a>
|
||
|
<a id="1009" class="Keyword">open</a> <a id="1014" href="Algebra.Construct.NaturalChoice.Base.html#990" class="Module">MinOperator</a> <a id="1026" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#797" class="Bound">minOp</a>
|
||
|
<a id="1032" class="Keyword">open</a> <a id="1037" href="Algebra.Construct.NaturalChoice.Base.html#1191" class="Module">MaxOperator</a> <a id="1049" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#823" class="Bound">maxOp</a>
|
||
|
|
||
|
<a id="1056" class="Keyword">open</a> <a id="1061" class="Keyword">import</a> <a id="1068" href="Algebra.Definitions.html" class="Module">Algebra.Definitions</a> <a id="1088" href="Relation.Binary.Bundles.html#2954" class="Function Operator">_≈_</a>
|
||
|
<a id="1092" class="Keyword">open</a> <a id="1097" class="Keyword">import</a> <a id="1104" href="Algebra.Structures.html" class="Module">Algebra.Structures</a> <a id="1123" href="Relation.Binary.Bundles.html#2954" class="Function Operator">_≈_</a>
|
||
|
<a id="1127" class="Keyword">open</a> <a id="1132" class="Keyword">import</a> <a id="1139" href="Algebra.Consequences.Setoid.html" class="Module">Algebra.Consequences.Setoid</a> <a id="1167" href="Relation.Binary.Bundles.html#2449" class="Function">Eq.setoid</a>
|
||
|
<a id="1177" class="Keyword">open</a> <a id="1182" class="Keyword">import</a> <a id="1189" href="Relation.Binary.Reasoning.Preorder.html" class="Module">Relation.Binary.Reasoning.Preorder</a> <a id="1224" href="Relation.Binary.Bundles.html#3197" class="Function">preorder</a>
|
||
|
|
||
|
<a id="1234" class="Comment">------------------------------------------------------------------------</a>
|
||
|
<a id="1307" class="Comment">-- Re-export properties of individual operators</a>
|
||
|
|
||
|
<a id="1356" class="Keyword">open</a> <a id="1361" class="Keyword">import</a> <a id="1368" href="Algebra.Construct.NaturalChoice.MinOp.html" class="Module">Algebra.Construct.NaturalChoice.MinOp</a> <a id="1406" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#797" class="Bound">minOp</a> <a id="1412" class="Keyword">public</a>
|
||
|
<a id="1419" class="Keyword">open</a> <a id="1424" class="Keyword">import</a> <a id="1431" href="Algebra.Construct.NaturalChoice.MaxOp.html" class="Module">Algebra.Construct.NaturalChoice.MaxOp</a> <a id="1469" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#823" class="Bound">maxOp</a> <a id="1475" class="Keyword">public</a>
|
||
|
|
||
|
<a id="1483" class="Comment">------------------------------------------------------------------------</a>
|
||
|
<a id="1556" class="Comment">-- Joint algebraic structures</a>
|
||
|
|
||
|
<a id="⊓-distribˡ-⊔"></a><a id="1587" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1587" class="Function">⊓-distribˡ-⊔</a> <a id="1600" class="Symbol">:</a> <a id="1602" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a> <a id="1606" href="Algebra.Definitions.html#3155" class="Function Operator">DistributesOverˡ</a> <a id="1623" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a>
|
||
|
<a id="1627" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1587" class="Function">⊓-distribˡ-⊔</a> <a id="1640" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1640" class="Bound">x</a> <a id="1642" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1642" class="Bound">y</a> <a id="1644" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1644" class="Bound">z</a> <a id="1646" class="Keyword">with</a> <a id="1651" href="Relation.Binary.Structures.html#3288" class="Function">total</a> <a id="1657" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1642" class="Bound">y</a> <a id="1659" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1644" class="Bound">z</a>
|
||
|
<a id="1661" class="Symbol">...</a> <a id="1665" class="Symbol">|</a> <a id="1667" href="Data.Sum.Base.html#675" class="InductiveConstructor">inj₁</a> <a id="1672" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1672" class="Bound">y≤z</a> <a id="1676" class="Symbol">=</a> <a id="1678" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="1695" class="Bound">x</a> <a id="1697" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="1699" class="Symbol">(</a><a id="1700" class="Bound">y</a> <a id="1702" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="1704" class="Bound">z</a><a id="1705" class="Symbol">)</a> <a id="1713" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="1717" href="Algebra.Construct.NaturalChoice.MinOp.html#1918" class="Function">⊓-congˡ</a> <a id="1725" class="Bound">x</a> <a id="1727" class="Symbol">(</a><a id="1728" href="Algebra.Construct.NaturalChoice.Base.html#1296" class="Field">x≤y⇒x⊔y≈y</a> <a id="1738" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1672" class="Bound">y≤z</a><a id="1741" class="Symbol">)</a> <a id="1743" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="1747" class="Bound">x</a> <a id="1749" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="1751" class="Bound">z</a> <a id="1765" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="1768" href="Algebra.Construct.NaturalChoice.Base.html#1296" class="Field">x≤y⇒x⊔y≈y</a> <a id="1778" class="Symbol">(</a><a id="1779" href="Algebra.Construct.NaturalChoice.MinOp.html#6966" class="Function">⊓-monoʳ-≤</a> <a id="1789" class="Bound">x</a> <a id="1791" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1672" class="Bound">y≤z</a><a id="1794" class="Symbol">)</a> <a id="1796" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
|
<a id="1800" class="Symbol">(</a><a id="1801" class="Bound">x</a> <a id="1803" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="1805" class="Bound">y</a><a id="1806" class="Symbol">)</a> <a id="1808" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="1810" class="Symbol">(</a><a id="1811" class="Bound">x</a> <a id="1813" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="1815" class="Bound">z</a><a id="1816" class="Symbol">)</a> <a id="1818" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
<a id="1820" class="Symbol">...</a> <a id="1824" class="Symbol">|</a> <a id="1826" href="Data.Sum.Base.html#700" class="InductiveConstructor">inj₂</a> <a id="1831" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1831" class="Bound">y≥z</a> <a id="1835" class="Symbol">=</a> <a id="1837" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="1854" class="Bound">x</a> <a id="1856" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="1858" class="Symbol">(</a><a id="1859" class="Bound">y</a> <a id="1861" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="1863" class="Bound">z</a><a id="1864" class="Symbol">)</a> <a id="1872" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="1876" href="Algebra.Construct.NaturalChoice.MinOp.html#1918" class="Function">⊓-congˡ</a> <a id="1884" class="Bound">x</a> <a id="1886" class="Symbol">(</a><a id="1887" href="Algebra.Construct.NaturalChoice.Base.html#1342" class="Field">x≥y⇒x⊔y≈x</a> <a id="1897" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1831" class="Bound">y≥z</a><a id="1900" class="Symbol">)</a> <a id="1902" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="1906" class="Bound">x</a> <a id="1908" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="1910" class="Bound">y</a> <a id="1924" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="1927" href="Algebra.Construct.NaturalChoice.Base.html#1342" class="Field">x≥y⇒x⊔y≈x</a> <a id="1937" class="Symbol">(</a><a id="1938" href="Algebra.Construct.NaturalChoice.MinOp.html#6966" class="Function">⊓-monoʳ-≤</a> <a id="1948" class="Bound">x</a> <a id="1950" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1831" class="Bound">y≥z</a><a id="1953" class="Symbol">)</a> <a id="1955" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
|
<a id="1959" class="Symbol">(</a><a id="1960" class="Bound">x</a> <a id="1962" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="1964" class="Bound">y</a><a id="1965" class="Symbol">)</a> <a id="1967" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="1969" class="Symbol">(</a><a id="1970" class="Bound">x</a> <a id="1972" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="1974" class="Bound">z</a><a id="1975" class="Symbol">)</a> <a id="1977" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
|
||
|
<a id="⊓-distribʳ-⊔"></a><a id="1980" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1980" class="Function">⊓-distribʳ-⊔</a> <a id="1993" class="Symbol">:</a> <a id="1995" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a> <a id="1999" href="Algebra.Definitions.html#3274" class="Function Operator">DistributesOverʳ</a> <a id="2016" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a>
|
||
|
<a id="2020" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1980" class="Function">⊓-distribʳ-⊔</a> <a id="2033" class="Symbol">=</a> <a id="2035" href="Algebra.Consequences.Setoid.html#15140" class="Function">comm+distrˡ⇒distrʳ</a> <a id="2054" href="Algebra.Construct.NaturalChoice.MaxOp.html#1034" class="Function">⊔-cong</a> <a id="2061" href="Algebra.Construct.NaturalChoice.MinOp.html#1732" class="Function">⊓-comm</a> <a id="2068" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1587" class="Function">⊓-distribˡ-⊔</a>
|
||
|
|
||
|
<a id="⊓-distrib-⊔"></a><a id="2082" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2082" class="Function">⊓-distrib-⊔</a> <a id="2094" class="Symbol">:</a> <a id="2096" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a> <a id="2100" href="Algebra.Definitions.html#3393" class="Function Operator">DistributesOver</a> <a id="2116" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a>
|
||
|
<a id="2120" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2082" class="Function">⊓-distrib-⊔</a> <a id="2132" class="Symbol">=</a> <a id="2134" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1587" class="Function">⊓-distribˡ-⊔</a> <a id="2147" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="2149" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#1980" class="Function">⊓-distribʳ-⊔</a>
|
||
|
|
||
|
<a id="⊔-distribˡ-⊓"></a><a id="2163" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2163" class="Function">⊔-distribˡ-⊓</a> <a id="2176" class="Symbol">:</a> <a id="2178" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a> <a id="2182" href="Algebra.Definitions.html#3155" class="Function Operator">DistributesOverˡ</a> <a id="2199" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a>
|
||
|
<a id="2203" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2163" class="Function">⊔-distribˡ-⊓</a> <a id="2216" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2216" class="Bound">x</a> <a id="2218" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2218" class="Bound">y</a> <a id="2220" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2220" class="Bound">z</a> <a id="2222" class="Keyword">with</a> <a id="2227" href="Relation.Binary.Structures.html#3288" class="Function">total</a> <a id="2233" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2218" class="Bound">y</a> <a id="2235" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2220" class="Bound">z</a>
|
||
|
<a id="2237" class="Symbol">...</a> <a id="2241" class="Symbol">|</a> <a id="2243" href="Data.Sum.Base.html#675" class="InductiveConstructor">inj₁</a> <a id="2248" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2248" class="Bound">y≤z</a> <a id="2252" class="Symbol">=</a> <a id="2254" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="2271" class="Bound">x</a> <a id="2273" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2275" class="Symbol">(</a><a id="2276" class="Bound">y</a> <a id="2278" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="2280" class="Bound">z</a><a id="2281" class="Symbol">)</a> <a id="2289" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2293" href="Algebra.Construct.NaturalChoice.MaxOp.html#1091" class="Function">⊔-congˡ</a> <a id="2301" class="Bound">x</a> <a id="2303" class="Symbol">(</a><a id="2304" href="Algebra.Construct.NaturalChoice.Base.html#1095" class="Function">x≤y⇒x⊓y≈x</a> <a id="2314" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2248" class="Bound">y≤z</a><a id="2317" class="Symbol">)</a> <a id="2319" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="2323" class="Bound">x</a> <a id="2325" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2327" class="Bound">y</a> <a id="2341" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="2344" href="Algebra.Construct.NaturalChoice.Base.html#1095" class="Function">x≤y⇒x⊓y≈x</a> <a id="2354" class="Symbol">(</a><a id="2355" href="Algebra.Construct.NaturalChoice.MaxOp.html#2362" class="Function">⊔-monoʳ-≤</a> <a id="2365" class="Bound">x</a> <a id="2367" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2248" class="Bound">y≤z</a><a id="2370" class="Symbol">)</a> <a id="2372" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
|
<a id="2376" class="Symbol">(</a><a id="2377" class="Bound">x</a> <a id="2379" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2381" class="Bound">y</a><a id="2382" class="Symbol">)</a> <a id="2384" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="2386" class="Symbol">(</a><a id="2387" class="Bound">x</a> <a id="2389" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2391" class="Bound">z</a><a id="2392" class="Symbol">)</a> <a id="2394" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
<a id="2396" class="Symbol">...</a> <a id="2400" class="Symbol">|</a> <a id="2402" href="Data.Sum.Base.html#700" class="InductiveConstructor">inj₂</a> <a id="2407" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2407" class="Bound">y≥z</a> <a id="2411" class="Symbol">=</a> <a id="2413" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="2430" class="Bound">x</a> <a id="2432" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2434" class="Symbol">(</a><a id="2435" class="Bound">y</a> <a id="2437" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="2439" class="Bound">z</a><a id="2440" class="Symbol">)</a> <a id="2448" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2452" href="Algebra.Construct.NaturalChoice.MaxOp.html#1091" class="Function">⊔-congˡ</a> <a id="2460" class="Bound">x</a> <a id="2462" class="Symbol">(</a><a id="2463" href="Algebra.Construct.NaturalChoice.Base.html#1141" class="Function">x≥y⇒x⊓y≈y</a> <a id="2473" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2407" class="Bound">y≥z</a><a id="2476" class="Symbol">)</a> <a id="2478" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="2482" class="Bound">x</a> <a id="2484" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2486" class="Bound">z</a> <a id="2500" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="2503" href="Algebra.Construct.NaturalChoice.Base.html#1141" class="Function">x≥y⇒x⊓y≈y</a> <a id="2513" class="Symbol">(</a><a id="2514" href="Algebra.Construct.NaturalChoice.MaxOp.html#2362" class="Function">⊔-monoʳ-≤</a> <a id="2524" class="Bound">x</a> <a id="2526" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2407" class="Bound">y≥z</a><a id="2529" class="Symbol">)</a> <a id="2531" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
|
<a id="2535" class="Symbol">(</a><a id="2536" class="Bound">x</a> <a id="2538" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2540" class="Bound">y</a><a id="2541" class="Symbol">)</a> <a id="2543" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="2545" class="Symbol">(</a><a id="2546" class="Bound">x</a> <a id="2548" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2550" class="Bound">z</a><a id="2551" class="Symbol">)</a> <a id="2553" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
|
||
|
<a id="⊔-distribʳ-⊓"></a><a id="2556" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2556" class="Function">⊔-distribʳ-⊓</a> <a id="2569" class="Symbol">:</a> <a id="2571" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a> <a id="2575" href="Algebra.Definitions.html#3274" class="Function Operator">DistributesOverʳ</a> <a id="2592" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a>
|
||
|
<a id="2596" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2556" class="Function">⊔-distribʳ-⊓</a> <a id="2609" class="Symbol">=</a> <a id="2611" href="Algebra.Consequences.Setoid.html#15140" class="Function">comm+distrˡ⇒distrʳ</a> <a id="2630" href="Algebra.Construct.NaturalChoice.MinOp.html#2418" class="Function">⊓-cong</a> <a id="2637" href="Algebra.Construct.NaturalChoice.MaxOp.html#1204" class="Function">⊔-comm</a> <a id="2644" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2163" class="Function">⊔-distribˡ-⊓</a>
|
||
|
|
||
|
<a id="⊔-distrib-⊓"></a><a id="2658" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2658" class="Function">⊔-distrib-⊓</a> <a id="2670" class="Symbol">:</a> <a id="2672" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a> <a id="2676" href="Algebra.Definitions.html#3393" class="Function Operator">DistributesOver</a> <a id="2692" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a>
|
||
|
<a id="2696" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2658" class="Function">⊔-distrib-⊓</a> <a id="2708" class="Symbol">=</a> <a id="2710" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2163" class="Function">⊔-distribˡ-⊓</a> <a id="2723" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="2725" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2556" class="Function">⊔-distribʳ-⊓</a>
|
||
|
|
||
|
<a id="⊓-absorbs-⊔"></a><a id="2739" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2739" class="Function">⊓-absorbs-⊔</a> <a id="2751" class="Symbol">:</a> <a id="2753" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a> <a id="2757" href="Algebra.Definitions.html#3918" class="Function Operator">Absorbs</a> <a id="2765" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a>
|
||
|
<a id="2769" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2739" class="Function">⊓-absorbs-⊔</a> <a id="2781" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2781" class="Bound">x</a> <a id="2783" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2783" class="Bound">y</a> <a id="2785" class="Keyword">with</a> <a id="2790" href="Relation.Binary.Structures.html#3288" class="Function">total</a> <a id="2796" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2781" class="Bound">x</a> <a id="2798" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2783" class="Bound">y</a>
|
||
|
<a id="2800" class="Symbol">...</a> <a id="2804" class="Symbol">|</a> <a id="2806" href="Data.Sum.Base.html#675" class="InductiveConstructor">inj₁</a> <a id="2811" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2811" class="Bound">x≤y</a> <a id="2815" class="Symbol">=</a> <a id="2817" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="2834" class="Bound">x</a> <a id="2836" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="2838" class="Symbol">(</a><a id="2839" class="Bound">x</a> <a id="2841" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2843" class="Bound">y</a><a id="2844" class="Symbol">)</a> <a id="2847" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2850" href="Algebra.Construct.NaturalChoice.MinOp.html#1918" class="Function">⊓-congˡ</a> <a id="2858" class="Bound">x</a> <a id="2860" class="Symbol">(</a><a id="2861" href="Algebra.Construct.NaturalChoice.Base.html#1296" class="Field">x≤y⇒x⊔y≈y</a> <a id="2871" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2811" class="Bound">x≤y</a><a id="2874" class="Symbol">)</a> <a id="2876" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="2880" class="Bound">x</a> <a id="2882" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="2884" class="Bound">y</a> <a id="2893" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2896" href="Algebra.Construct.NaturalChoice.Base.html#1095" class="Function">x≤y⇒x⊓y≈x</a> <a id="2906" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2811" class="Bound">x≤y</a> <a id="2910" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="2914" class="Bound">x</a> <a id="2927" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
<a id="2929" class="Symbol">...</a> <a id="2933" class="Symbol">|</a> <a id="2935" href="Data.Sum.Base.html#700" class="InductiveConstructor">inj₂</a> <a id="2940" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2940" class="Bound">y≤x</a> <a id="2944" class="Symbol">=</a> <a id="2946" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="2963" class="Bound">x</a> <a id="2965" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="2967" class="Symbol">(</a><a id="2968" class="Bound">x</a> <a id="2970" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="2972" class="Bound">y</a><a id="2973" class="Symbol">)</a> <a id="2976" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2979" href="Algebra.Construct.NaturalChoice.MinOp.html#1918" class="Function">⊓-congˡ</a> <a id="2987" class="Bound">x</a> <a id="2989" class="Symbol">(</a><a id="2990" href="Algebra.Construct.NaturalChoice.Base.html#1342" class="Field">x≥y⇒x⊔y≈x</a> <a id="3000" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2940" class="Bound">y≤x</a><a id="3003" class="Symbol">)</a> <a id="3005" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="3009" class="Bound">x</a> <a id="3011" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="3013" class="Bound">x</a> <a id="3022" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3025" href="Algebra.Construct.NaturalChoice.MinOp.html#3166" class="Function">⊓-idem</a> <a id="3032" class="Bound">x</a> <a id="3034" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="3038" class="Bound">x</a> <a id="3051" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
|
||
|
<a id="⊔-absorbs-⊓"></a><a id="3054" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3054" class="Function">⊔-absorbs-⊓</a> <a id="3066" class="Symbol">:</a> <a id="3068" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a> <a id="3072" href="Algebra.Definitions.html#3918" class="Function Operator">Absorbs</a> <a id="3080" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a>
|
||
|
<a id="3084" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3054" class="Function">⊔-absorbs-⊓</a> <a id="3096" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3096" class="Bound">x</a> <a id="3098" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3098" class="Bound">y</a> <a id="3100" class="Keyword">with</a> <a id="3105" href="Relation.Binary.Structures.html#3288" class="Function">total</a> <a id="3111" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3096" class="Bound">x</a> <a id="3113" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3098" class="Bound">y</a>
|
||
|
<a id="3115" class="Symbol">...</a> <a id="3119" class="Symbol">|</a> <a id="3121" href="Data.Sum.Base.html#675" class="InductiveConstructor">inj₁</a> <a id="3126" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3126" class="Bound">x≤y</a> <a id="3130" class="Symbol">=</a> <a id="3132" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="3149" class="Bound">x</a> <a id="3151" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="3153" class="Symbol">(</a><a id="3154" class="Bound">x</a> <a id="3156" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="3158" class="Bound">y</a><a id="3159" class="Symbol">)</a> <a id="3162" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3165" href="Algebra.Construct.NaturalChoice.MaxOp.html#1091" class="Function">⊔-congˡ</a> <a id="3173" class="Bound">x</a> <a id="3175" class="Symbol">(</a><a id="3176" href="Algebra.Construct.NaturalChoice.Base.html#1095" class="Function">x≤y⇒x⊓y≈x</a> <a id="3186" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3126" class="Bound">x≤y</a><a id="3189" class="Symbol">)</a> <a id="3191" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="3195" class="Bound">x</a> <a id="3197" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="3199" class="Bound">x</a> <a id="3208" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3211" href="Algebra.Construct.NaturalChoice.MaxOp.html#1120" class="Function">⊔-idem</a> <a id="3218" class="Bound">x</a> <a id="3220" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="3224" class="Bound">x</a> <a id="3237" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
<a id="3239" class="Symbol">...</a> <a id="3243" class="Symbol">|</a> <a id="3245" href="Data.Sum.Base.html#700" class="InductiveConstructor">inj₂</a> <a id="3250" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3250" class="Bound">y≤x</a> <a id="3254" class="Symbol">=</a> <a id="3256" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="3273" class="Bound">x</a> <a id="3275" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="3277" class="Symbol">(</a><a id="3278" class="Bound">x</a> <a id="3280" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="3282" class="Bound">y</a><a id="3283" class="Symbol">)</a> <a id="3286" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3289" href="Algebra.Construct.NaturalChoice.MaxOp.html#1091" class="Function">⊔-congˡ</a> <a id="3297" class="Bound">x</a> <a id="3299" class="Symbol">(</a><a id="3300" href="Algebra.Construct.NaturalChoice.Base.html#1141" class="Function">x≥y⇒x⊓y≈y</a> <a id="3310" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3250" class="Bound">y≤x</a><a id="3313" class="Symbol">)</a> <a id="3315" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="3319" class="Bound">x</a> <a id="3321" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="3323" class="Bound">y</a> <a id="3332" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3335" href="Algebra.Construct.NaturalChoice.Base.html#1342" class="Field">x≥y⇒x⊔y≈x</a> <a id="3345" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3250" class="Bound">y≤x</a> <a id="3349" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="3353" class="Bound">x</a> <a id="3366" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
|
||
|
<a id="⊔-⊓-absorptive"></a><a id="3369" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3369" class="Function">⊔-⊓-absorptive</a> <a id="3384" class="Symbol">:</a> <a id="3386" href="Algebra.Definitions.html#3997" class="Function">Absorptive</a> <a id="3397" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a> <a id="3401" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a>
|
||
|
<a id="3405" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3369" class="Function">⊔-⊓-absorptive</a> <a id="3420" class="Symbol">=</a> <a id="3422" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3054" class="Function">⊔-absorbs-⊓</a> <a id="3434" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="3436" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2739" class="Function">⊓-absorbs-⊔</a>
|
||
|
|
||
|
<a id="⊓-⊔-absorptive"></a><a id="3449" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3449" class="Function">⊓-⊔-absorptive</a> <a id="3464" class="Symbol">:</a> <a id="3466" href="Algebra.Definitions.html#3997" class="Function">Absorptive</a> <a id="3477" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">_⊓_</a> <a id="3481" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">_⊔_</a>
|
||
|
<a id="3485" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3449" class="Function">⊓-⊔-absorptive</a> <a id="3500" class="Symbol">=</a> <a id="3502" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#2739" class="Function">⊓-absorbs-⊔</a> <a id="3514" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="3516" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3054" class="Function">⊔-absorbs-⊓</a>
|
||
|
|
||
|
<a id="3529" class="Comment">------------------------------------------------------------------------</a>
|
||
|
<a id="3602" class="Comment">-- Other joint properties</a>
|
||
|
|
||
|
<a id="3629" class="Keyword">private</a> <a id="_≥_"></a><a id="3637" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3637" class="Function Operator">_≥_</a> <a id="3641" class="Symbol">=</a> <a id="3643" href="Function.Base.html#1638" class="Function">flip</a> <a id="3648" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#921" class="Function Operator">_≤_</a>
|
||
|
|
||
|
<a id="antimono-≤-distrib-⊓"></a><a id="3653" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3653" class="Function">antimono-≤-distrib-⊓</a> <a id="3674" class="Symbol">:</a> <a id="3676" class="Symbol">∀</a> <a id="3678" class="Symbol">{</a><a id="3679" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3679" class="Bound">f</a><a id="3680" class="Symbol">}</a> <a id="3682" class="Symbol">→</a> <a id="3684" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3679" class="Bound">f</a> <a id="3686" href="Relation.Binary.Core.html#1577" class="Function Operator">Preserves</a> <a id="3696" href="Relation.Binary.Bundles.html#2954" class="Function Operator">_≈_</a> <a id="3700" href="Relation.Binary.Core.html#1577" class="Function Operator">⟶</a> <a id="3702" href="Relation.Binary.Bundles.html#2954" class="Function Operator">_≈_</a> <a id="3706" class="Symbol">→</a> <a id="3708" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3679" class="Bound">f</a> <a id="3710" href="Relation.Binary.Core.html#1577" class="Function Operator">Preserves</a> <a id="3720" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#921" class="Function Operator">_≤_</a> <a id="3724" href="Relation.Binary.Core.html#1577" class="Function Operator">⟶</a> <a id="3726" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3637" class="Function Operator">_≥_</a> <a id="3730" class="Symbol">→</a>
|
||
|
<a id="3755" class="Symbol">∀</a> <a id="3757" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3757" class="Bound">x</a> <a id="3759" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3759" class="Bound">y</a> <a id="3761" class="Symbol">→</a> <a id="3763" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3679" class="Bound">f</a> <a id="3765" class="Symbol">(</a><a id="3766" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3757" class="Bound">x</a> <a id="3768" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="3770" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3759" class="Bound">y</a><a id="3771" class="Symbol">)</a> <a id="3773" href="Relation.Binary.Bundles.html#2954" class="Function Operator">≈</a> <a id="3775" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3679" class="Bound">f</a> <a id="3777" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3757" class="Bound">x</a> <a id="3779" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="3781" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3679" class="Bound">f</a> <a id="3783" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3759" class="Bound">y</a>
|
||
|
<a id="3785" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3653" class="Function">antimono-≤-distrib-⊓</a> <a id="3806" class="Symbol">{</a><a id="3807" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3807" class="Bound">f</a><a id="3808" class="Symbol">}</a> <a id="3810" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3810" class="Bound">cong</a> <a id="3815" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3815" class="Bound">antimono</a> <a id="3824" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3824" class="Bound">x</a> <a id="3826" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3826" class="Bound">y</a> <a id="3828" class="Keyword">with</a> <a id="3833" href="Relation.Binary.Structures.html#3288" class="Function">total</a> <a id="3839" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3824" class="Bound">x</a> <a id="3841" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3826" class="Bound">y</a>
|
||
|
<a id="3843" class="Symbol">...</a> <a id="3847" class="Symbol">|</a> <a id="3849" href="Data.Sum.Base.html#675" class="InductiveConstructor">inj₁</a> <a id="3854" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3854" class="Bound">x≤y</a> <a id="3858" class="Symbol">=</a> <a id="3860" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="3877" class="Bound">f</a> <a id="3879" class="Symbol">(</a><a id="3880" class="Bound">x</a> <a id="3882" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="3884" class="Bound">y</a><a id="3885" class="Symbol">)</a> <a id="3888" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="3891" class="Bound">cong</a> <a id="3896" class="Symbol">(</a><a id="3897" href="Algebra.Construct.NaturalChoice.Base.html#1095" class="Function">x≤y⇒x⊓y≈x</a> <a id="3907" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3854" class="Bound">x≤y</a><a id="3910" class="Symbol">)</a> <a id="3912" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="3916" class="Bound">f</a> <a id="3918" class="Bound">x</a> <a id="3927" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="3930" href="Algebra.Construct.NaturalChoice.Base.html#1342" class="Field">x≥y⇒x⊔y≈x</a> <a id="3940" class="Symbol">(</a><a id="3941" class="Bound">antimono</a> <a id="3950" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3854" class="Bound">x≤y</a><a id="3953" class="Symbol">)</a> <a id="3955" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
|
<a id="3959" class="Bound">f</a> <a id="3961" class="Bound">x</a> <a id="3963" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="3965" class="Bound">f</a> <a id="3967" class="Bound">y</a> <a id="3970" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
<a id="3972" class="Symbol">...</a> <a id="3976" class="Symbol">|</a> <a id="3978" href="Data.Sum.Base.html#700" class="InductiveConstructor">inj₂</a> <a id="3983" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3983" class="Bound">y≤x</a> <a id="3987" class="Symbol">=</a> <a id="3989" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="4006" class="Bound">f</a> <a id="4008" class="Symbol">(</a><a id="4009" class="Bound">x</a> <a id="4011" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="4013" class="Bound">y</a><a id="4014" class="Symbol">)</a> <a id="4017" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4020" class="Bound">cong</a> <a id="4025" class="Symbol">(</a><a id="4026" href="Algebra.Construct.NaturalChoice.Base.html#1141" class="Function">x≥y⇒x⊓y≈y</a> <a id="4036" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3983" class="Bound">y≤x</a><a id="4039" class="Symbol">)</a> <a id="4041" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="4045" class="Bound">f</a> <a id="4047" class="Bound">y</a> <a id="4056" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="4059" href="Algebra.Construct.NaturalChoice.Base.html#1296" class="Field">x≤y⇒x⊔y≈y</a> <a id="4069" class="Symbol">(</a><a id="4070" class="Bound">antimono</a> <a id="4079" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3983" class="Bound">y≤x</a><a id="4082" class="Symbol">)</a> <a id="4084" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
|
<a id="4088" class="Bound">f</a> <a id="4090" class="Bound">x</a> <a id="4092" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="4094" class="Bound">f</a> <a id="4096" class="Bound">y</a> <a id="4099" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
|
||
|
<a id="antimono-≤-distrib-⊔"></a><a id="4102" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4102" class="Function">antimono-≤-distrib-⊔</a> <a id="4123" class="Symbol">:</a> <a id="4125" class="Symbol">∀</a> <a id="4127" class="Symbol">{</a><a id="4128" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4128" class="Bound">f</a><a id="4129" class="Symbol">}</a> <a id="4131" class="Symbol">→</a> <a id="4133" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4128" class="Bound">f</a> <a id="4135" href="Relation.Binary.Core.html#1577" class="Function Operator">Preserves</a> <a id="4145" href="Relation.Binary.Bundles.html#2954" class="Function Operator">_≈_</a> <a id="4149" href="Relation.Binary.Core.html#1577" class="Function Operator">⟶</a> <a id="4151" href="Relation.Binary.Bundles.html#2954" class="Function Operator">_≈_</a> <a id="4155" class="Symbol">→</a> <a id="4157" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4128" class="Bound">f</a> <a id="4159" href="Relation.Binary.Core.html#1577" class="Function Operator">Preserves</a> <a id="4169" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#921" class="Function Operator">_≤_</a> <a id="4173" href="Relation.Binary.Core.html#1577" class="Function Operator">⟶</a> <a id="4175" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#3637" class="Function Operator">_≥_</a> <a id="4179" class="Symbol">→</a>
|
||
|
<a id="4204" class="Symbol">∀</a> <a id="4206" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4206" class="Bound">x</a> <a id="4208" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4208" class="Bound">y</a> <a id="4210" class="Symbol">→</a> <a id="4212" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4128" class="Bound">f</a> <a id="4214" class="Symbol">(</a><a id="4215" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4206" class="Bound">x</a> <a id="4217" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="4219" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4208" class="Bound">y</a><a id="4220" class="Symbol">)</a> <a id="4222" href="Relation.Binary.Bundles.html#2954" class="Function Operator">≈</a> <a id="4224" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4128" class="Bound">f</a> <a id="4226" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4206" class="Bound">x</a> <a id="4228" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="4230" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4128" class="Bound">f</a> <a id="4232" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4208" class="Bound">y</a>
|
||
|
<a id="4234" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4102" class="Function">antimono-≤-distrib-⊔</a> <a id="4255" class="Symbol">{</a><a id="4256" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4256" class="Bound">f</a><a id="4257" class="Symbol">}</a> <a id="4259" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4259" class="Bound">cong</a> <a id="4264" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4264" class="Bound">antimono</a> <a id="4273" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4273" class="Bound">x</a> <a id="4275" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4275" class="Bound">y</a> <a id="4277" class="Keyword">with</a> <a id="4282" href="Relation.Binary.Structures.html#3288" class="Function">total</a> <a id="4288" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4273" class="Bound">x</a> <a id="4290" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4275" class="Bound">y</a>
|
||
|
<a id="4292" class="Symbol">...</a> <a id="4296" class="Symbol">|</a> <a id="4298" href="Data.Sum.Base.html#675" class="InductiveConstructor">inj₁</a> <a id="4303" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4303" class="Bound">x≤y</a> <a id="4307" class="Symbol">=</a> <a id="4309" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="4326" class="Bound">f</a> <a id="4328" class="Symbol">(</a><a id="4329" class="Bound">x</a> <a id="4331" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="4333" class="Bound">y</a><a id="4334" class="Symbol">)</a> <a id="4337" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4340" class="Bound">cong</a> <a id="4345" class="Symbol">(</a><a id="4346" href="Algebra.Construct.NaturalChoice.Base.html#1296" class="Field">x≤y⇒x⊔y≈y</a> <a id="4356" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4303" class="Bound">x≤y</a><a id="4359" class="Symbol">)</a> <a id="4361" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="4365" class="Bound">f</a> <a id="4367" class="Bound">y</a> <a id="4376" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="4379" href="Algebra.Construct.NaturalChoice.Base.html#1141" class="Function">x≥y⇒x⊓y≈y</a> <a id="4389" class="Symbol">(</a><a id="4390" class="Bound">antimono</a> <a id="4399" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4303" class="Bound">x≤y</a><a id="4402" class="Symbol">)</a> <a id="4404" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
|
<a id="4408" class="Bound">f</a> <a id="4410" class="Bound">x</a> <a id="4412" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="4414" class="Bound">f</a> <a id="4416" class="Bound">y</a> <a id="4419" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
<a id="4421" class="Symbol">...</a> <a id="4425" class="Symbol">|</a> <a id="4427" href="Data.Sum.Base.html#700" class="InductiveConstructor">inj₂</a> <a id="4432" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4432" class="Bound">y≤x</a> <a id="4436" class="Symbol">=</a> <a id="4438" href="Relation.Binary.Reasoning.Syntax.html#2557" class="Function Operator">begin-equality</a>
|
||
|
<a id="4455" class="Bound">f</a> <a id="4457" class="Symbol">(</a><a id="4458" class="Bound">x</a> <a id="4460" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="4462" class="Bound">y</a><a id="4463" class="Symbol">)</a> <a id="4466" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4469" class="Bound">cong</a> <a id="4474" class="Symbol">(</a><a id="4475" href="Algebra.Construct.NaturalChoice.Base.html#1342" class="Field">x≥y⇒x⊔y≈x</a> <a id="4485" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4432" class="Bound">y≤x</a><a id="4488" class="Symbol">)</a> <a id="4490" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
|
<a id="4494" class="Bound">f</a> <a id="4496" class="Bound">x</a> <a id="4505" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="4508" href="Algebra.Construct.NaturalChoice.Base.html#1095" class="Function">x≤y⇒x⊓y≈x</a> <a id="4518" class="Symbol">(</a><a id="4519" class="Bound">antimono</a> <a id="4528" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4432" class="Bound">y≤x</a><a id="4531" class="Symbol">)</a> <a id="4533" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
|
<a id="4537" class="Bound">f</a> <a id="4539" class="Bound">x</a> <a id="4541" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="4543" class="Bound">f</a> <a id="4545" class="Bound">y</a> <a id="4548" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
|
||
|
<a id="x⊓y≤x⊔y"></a><a id="4551" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4551" class="Function">x⊓y≤x⊔y</a> <a id="4559" class="Symbol">:</a> <a id="4561" class="Symbol">∀</a> <a id="4563" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4563" class="Bound">x</a> <a id="4565" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4565" class="Bound">y</a> <a id="4567" class="Symbol">→</a> <a id="4569" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4563" class="Bound">x</a> <a id="4571" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="4573" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4565" class="Bound">y</a> <a id="4575" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#921" class="Function Operator">≤</a> <a id="4577" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4563" class="Bound">x</a> <a id="4579" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="4581" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4565" class="Bound">y</a>
|
||
|
<a id="4583" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4551" class="Function">x⊓y≤x⊔y</a> <a id="4591" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4591" class="Bound">x</a> <a id="4593" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4593" class="Bound">y</a> <a id="4595" class="Symbol">=</a> <a id="4597" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
|
<a id="4605" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4591" class="Bound">x</a> <a id="4607" href="Algebra.Construct.NaturalChoice.Base.html#1065" class="Function Operator">⊓</a> <a id="4609" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4593" class="Bound">y</a> <a id="4611" href="Relation.Binary.Reasoning.Syntax.html#5211" class="Function">∼⟨</a> <a id="4614" href="Algebra.Construct.NaturalChoice.MinOp.html#1304" class="Function">x⊓y≤x</a> <a id="4620" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4591" class="Bound">x</a> <a id="4622" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4593" class="Bound">y</a> <a id="4624" href="Relation.Binary.Reasoning.Syntax.html#5211" class="Function">⟩</a>
|
||
|
<a id="4628" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4591" class="Bound">x</a> <a id="4634" href="Relation.Binary.Reasoning.Syntax.html#5211" class="Function">∼⟨</a> <a id="4637" href="Algebra.Construct.NaturalChoice.MaxOp.html#2045" class="Function">x≤x⊔y</a> <a id="4643" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4591" class="Bound">x</a> <a id="4645" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4593" class="Bound">y</a> <a id="4647" href="Relation.Binary.Reasoning.Syntax.html#5211" class="Function">⟩</a>
|
||
|
<a id="4651" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4591" class="Bound">x</a> <a id="4653" href="Algebra.Construct.NaturalChoice.Base.html#1266" class="Field Operator">⊔</a> <a id="4655" href="Algebra.Construct.NaturalChoice.MinMaxOp.html#4593" class="Bound">y</a> <a id="4657" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
</pre></body></html>
|