mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
540 lines
No EOL
291 KiB
HTML
540 lines
No EOL
291 KiB
HTML
<!DOCTYPE HTML>
|
||
<html><head><meta charset="utf-8"><title>Algebra.Lattice.Properties.BooleanAlgebra</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 derivable properties of Boolean algebras</a>
|
||
<a id="155" class="Comment">------------------------------------------------------------------------</a>
|
||
|
||
<a id="229" class="Symbol">{-#</a> <a id="233" class="Keyword">OPTIONS</a> <a id="241" class="Pragma">--cubical-compatible</a> <a id="262" class="Pragma">--safe</a> <a id="269" class="Symbol">#-}</a>
|
||
|
||
<a id="274" class="Keyword">open</a> <a id="279" class="Keyword">import</a> <a id="286" href="Algebra.Lattice.Bundles.html" class="Module">Algebra.Lattice.Bundles</a>
|
||
|
||
<a id="311" class="Keyword">module</a> <a id="318" href="Algebra.Lattice.Properties.BooleanAlgebra.html" class="Module">Algebra.Lattice.Properties.BooleanAlgebra</a>
|
||
<a id="362" class="Symbol">{</a><a id="363" href="Algebra.Lattice.Properties.BooleanAlgebra.html#363" class="Bound">b₁</a> <a id="366" href="Algebra.Lattice.Properties.BooleanAlgebra.html#366" class="Bound">b₂</a><a id="368" class="Symbol">}</a> <a id="370" class="Symbol">(</a><a id="371" href="Algebra.Lattice.Properties.BooleanAlgebra.html#371" class="Bound">B</a> <a id="373" class="Symbol">:</a> <a id="375" href="Algebra.Lattice.Bundles.html#5554" class="Record">BooleanAlgebra</a> <a id="390" href="Algebra.Lattice.Properties.BooleanAlgebra.html#363" class="Bound">b₁</a> <a id="393" href="Algebra.Lattice.Properties.BooleanAlgebra.html#366" class="Bound">b₂</a><a id="395" class="Symbol">)</a>
|
||
<a id="399" class="Keyword">where</a>
|
||
|
||
<a id="406" class="Keyword">open</a> <a id="411" href="Algebra.Lattice.Bundles.html#5554" class="Module">BooleanAlgebra</a> <a id="426" href="Algebra.Lattice.Properties.BooleanAlgebra.html#371" class="Bound">B</a>
|
||
|
||
<a id="429" class="Keyword">import</a> <a id="436" href="Algebra.Lattice.Properties.DistributiveLattice.html" class="Module">Algebra.Lattice.Properties.DistributiveLattice</a> <a id="483" class="Symbol">as</a> <a id="486" class="Module">DistribLatticeProperties</a>
|
||
<a id="511" class="Keyword">open</a> <a id="516" class="Keyword">import</a> <a id="523" href="Algebra.Core.html" class="Module">Algebra.Core</a>
|
||
<a id="536" class="Keyword">open</a> <a id="541" class="Keyword">import</a> <a id="548" href="Algebra.Structures.html" class="Module">Algebra.Structures</a> <a id="567" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">_≈_</a>
|
||
<a id="571" class="Keyword">open</a> <a id="576" class="Keyword">import</a> <a id="583" href="Algebra.Definitions.html" class="Module">Algebra.Definitions</a> <a id="603" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">_≈_</a>
|
||
<a id="607" class="Keyword">open</a> <a id="612" class="Keyword">import</a> <a id="619" href="Algebra.Consequences.Setoid.html" class="Module">Algebra.Consequences.Setoid</a> <a id="647" href="Algebra.Lattice.Bundles.html#4863" class="Function">setoid</a>
|
||
<a id="654" class="Keyword">open</a> <a id="659" class="Keyword">import</a> <a id="666" href="Algebra.Bundles.html" class="Module">Algebra.Bundles</a>
|
||
<a id="682" class="Keyword">open</a> <a id="687" class="Keyword">import</a> <a id="694" href="Algebra.Lattice.Structures.html" class="Module">Algebra.Lattice.Structures</a> <a id="721" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">_≈_</a>
|
||
<a id="725" class="Keyword">open</a> <a id="730" class="Keyword">import</a> <a id="737" href="Relation.Binary.Reasoning.Setoid.html" class="Module">Relation.Binary.Reasoning.Setoid</a> <a id="770" href="Algebra.Lattice.Bundles.html#4863" class="Function">setoid</a>
|
||
<a id="777" class="Keyword">open</a> <a id="782" class="Keyword">import</a> <a id="789" href="Function.Base.html" class="Module">Function.Base</a> <a id="803" class="Keyword">using</a> <a id="809" class="Symbol">(</a><a id="810" href="Function.Base.html#704" class="Function">id</a><a id="812" class="Symbol">;</a> <a id="814" href="Function.Base.html#1974" class="Function Operator">_$_</a><a id="817" class="Symbol">;</a> <a id="819" href="Function.Base.html#4322" class="Function Operator">_⟨_⟩_</a><a id="824" class="Symbol">)</a>
|
||
<a id="826" class="Keyword">open</a> <a id="831" class="Keyword">import</a> <a id="838" href="Function.Bundles.html" class="Module">Function.Bundles</a> <a id="855" class="Keyword">using</a> <a id="861" class="Symbol">(</a><a id="862" href="Function.Bundles.html#12039" class="Function Operator">_⇔_</a><a id="865" class="Symbol">;</a> <a id="867" class="Keyword">module</a> <a id="874" href="Function.Bundles.html#4752" class="Module">Equivalence</a><a id="885" class="Symbol">)</a>
|
||
<a id="887" class="Keyword">open</a> <a id="892" class="Keyword">import</a> <a id="899" href="Data.Product.Base.html" class="Module">Data.Product.Base</a> <a id="917" class="Keyword">using</a> <a id="923" class="Symbol">(</a><a id="924" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">_,_</a><a id="927" class="Symbol">)</a>
|
||
|
||
<a id="930" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="1003" class="Comment">-- Export properties from distributive lattices</a>
|
||
|
||
<a id="1052" class="Keyword">open</a> <a id="1057" href="Algebra.Lattice.Properties.DistributiveLattice.html" class="Module">DistribLatticeProperties</a> <a id="1082" href="Algebra.Lattice.Bundles.html#6010" class="Function">distributiveLattice</a> <a id="1102" class="Keyword">public</a>
|
||
|
||
<a id="1110" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="1183" class="Comment">-- The dual construction is also a boolean algebra</a>
|
||
|
||
<a id="∧-∨-isBooleanAlgebra"></a><a id="1235" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1235" class="Function">∧-∨-isBooleanAlgebra</a> <a id="1256" class="Symbol">:</a> <a id="1258" href="Algebra.Lattice.Structures.html#5413" class="Record">IsBooleanAlgebra</a> <a id="1275" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="1279" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a> <a id="1283" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬_</a> <a id="1286" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="1288" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a>
|
||
<a id="1290" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1235" class="Function">∧-∨-isBooleanAlgebra</a> <a id="1311" class="Symbol">=</a> <a id="1313" class="Keyword">record</a>
|
||
<a id="1322" class="Symbol">{</a> <a id="1324" href="Algebra.Lattice.Structures.html#5501" class="Field">isDistributiveLattice</a> <a id="1346" class="Symbol">=</a> <a id="1348" href="Algebra.Lattice.Properties.DistributiveLattice.html#901" class="Function">∧-∨-isDistributiveLattice</a>
|
||
<a id="1376" class="Symbol">;</a> <a id="1378" href="Algebra.Lattice.Structures.html#5555" class="Field">∨-complement</a> <a id="1400" class="Symbol">=</a> <a id="1402" href="Algebra.Lattice.Structures.html#5597" class="Function">∧-complement</a>
|
||
<a id="1417" class="Symbol">;</a> <a id="1419" href="Algebra.Lattice.Structures.html#5597" class="Field">∧-complement</a> <a id="1441" class="Symbol">=</a> <a id="1443" href="Algebra.Lattice.Structures.html#5555" class="Function">∨-complement</a>
|
||
<a id="1458" class="Symbol">;</a> <a id="1460" href="Algebra.Lattice.Structures.html#5639" class="Field">¬-cong</a> <a id="1482" class="Symbol">=</a> <a id="1484" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a>
|
||
<a id="1493" class="Symbol">}</a>
|
||
|
||
<a id="∧-∨-booleanAlgebra"></a><a id="1496" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1496" class="Function">∧-∨-booleanAlgebra</a> <a id="1515" class="Symbol">:</a> <a id="1517" href="Algebra.Lattice.Bundles.html#5554" class="Record">BooleanAlgebra</a> <a id="1532" class="Symbol">_</a> <a id="1534" class="Symbol">_</a>
|
||
<a id="1536" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1496" class="Function">∧-∨-booleanAlgebra</a> <a id="1555" class="Symbol">=</a> <a id="1557" class="Keyword">record</a>
|
||
<a id="1566" class="Symbol">{</a> <a id="1568" href="Algebra.Lattice.Bundles.html#5903" class="Field">isBooleanAlgebra</a> <a id="1585" class="Symbol">=</a> <a id="1587" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1235" class="Function">∧-∨-isBooleanAlgebra</a>
|
||
<a id="1610" class="Symbol">}</a>
|
||
|
||
<a id="1613" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="1686" class="Comment">-- (∨, ∧, ⊥, ⊤) and (∧, ∨, ⊤, ⊥) are commutative semirings</a>
|
||
|
||
<a id="∧-identityʳ"></a><a id="1746" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1746" class="Function">∧-identityʳ</a> <a id="1758" class="Symbol">:</a> <a id="1760" href="Algebra.Definitions.html#1781" class="Function">RightIdentity</a> <a id="1774" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="1776" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a>
|
||
<a id="1780" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1746" class="Function">∧-identityʳ</a> <a id="1792" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1792" class="Bound">x</a> <a id="1794" class="Symbol">=</a> <a id="1796" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="1804" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1792" class="Bound">x</a> <a id="1806" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="1808" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="1819" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="1822" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="1830" class="Symbol">(</a><a id="1831" href="Relation.Binary.Structures.html#1622" class="Function">sym</a> <a id="1835" class="Symbol">(</a><a id="1836" href="Algebra.Lattice.Structures.html#5812" class="Function">∨-complementʳ</a> <a id="1850" class="Symbol">_))</a> <a id="1854" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="1858" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1792" class="Bound">x</a> <a id="1860" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="1862" class="Symbol">(</a><a id="1863" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1792" class="Bound">x</a> <a id="1865" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="1867" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="1869" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1792" class="Bound">x</a><a id="1870" class="Symbol">)</a> <a id="1873" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="1876" href="Algebra.Lattice.Structures.html#4454" class="Function">∧-absorbs-∨</a> <a id="1888" class="Symbol">_</a> <a id="1890" class="Symbol">_</a> <a id="1892" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="1896" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1792" class="Bound">x</a> <a id="1911" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="∧-identityˡ"></a><a id="1914" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1914" class="Function">∧-identityˡ</a> <a id="1926" class="Symbol">:</a> <a id="1928" href="Algebra.Definitions.html#1708" class="Function">LeftIdentity</a> <a id="1941" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="1943" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a>
|
||
<a id="1947" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1914" class="Function">∧-identityˡ</a> <a id="1959" class="Symbol">=</a> <a id="1961" href="Algebra.Consequences.Setoid.html#4478" class="Function">comm∧idʳ⇒idˡ</a> <a id="1974" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="1981" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1746" class="Function">∧-identityʳ</a>
|
||
|
||
<a id="∧-identity"></a><a id="1994" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1994" class="Function">∧-identity</a> <a id="2005" class="Symbol">:</a> <a id="2007" href="Algebra.Definitions.html#1856" class="Function">Identity</a> <a id="2016" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="2018" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a>
|
||
<a id="2022" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1994" class="Function">∧-identity</a> <a id="2033" class="Symbol">=</a> <a id="2035" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1914" class="Function">∧-identityˡ</a> <a id="2047" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="2049" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1746" class="Function">∧-identityʳ</a>
|
||
|
||
<a id="∨-identityʳ"></a><a id="2062" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2062" class="Function">∨-identityʳ</a> <a id="2074" class="Symbol">:</a> <a id="2076" href="Algebra.Definitions.html#1781" class="Function">RightIdentity</a> <a id="2090" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="2092" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a>
|
||
<a id="2096" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2062" class="Function">∨-identityʳ</a> <a id="2108" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2108" class="Bound">x</a> <a id="2110" class="Symbol">=</a> <a id="2112" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="2120" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2108" class="Bound">x</a> <a id="2122" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="2124" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="2135" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2138" href="Algebra.Lattice.Structures.html#4639" class="Function">∨-congˡ</a> <a id="2146" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="2148" href="Relation.Binary.Structures.html#1622" class="Function">sym</a> <a id="2152" class="Symbol">(</a><a id="2153" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="2167" class="Symbol">_)</a> <a id="2170" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2174" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2108" class="Bound">x</a> <a id="2176" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="2178" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2108" class="Bound">x</a> <a id="2180" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="2182" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="2184" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2108" class="Bound">x</a> <a id="2189" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2192" href="Algebra.Lattice.Structures.html#4392" class="Function">∨-absorbs-∧</a> <a id="2204" class="Symbol">_</a> <a id="2206" class="Symbol">_</a> <a id="2208" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2212" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2108" class="Bound">x</a> <a id="2227" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="∨-identityˡ"></a><a id="2230" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2230" class="Function">∨-identityˡ</a> <a id="2242" class="Symbol">:</a> <a id="2244" href="Algebra.Definitions.html#1708" class="Function">LeftIdentity</a> <a id="2257" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="2259" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a>
|
||
<a id="2263" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2230" class="Function">∨-identityˡ</a> <a id="2275" class="Symbol">=</a> <a id="2277" href="Algebra.Consequences.Setoid.html#4478" class="Function">comm∧idʳ⇒idˡ</a> <a id="2290" href="Algebra.Lattice.Structures.html#4113" class="Function">∨-comm</a> <a id="2297" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2062" class="Function">∨-identityʳ</a>
|
||
|
||
<a id="∨-identity"></a><a id="2310" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2310" class="Function">∨-identity</a> <a id="2321" class="Symbol">:</a> <a id="2323" href="Algebra.Definitions.html#1856" class="Function">Identity</a> <a id="2332" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="2334" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a>
|
||
<a id="2338" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2310" class="Function">∨-identity</a> <a id="2349" class="Symbol">=</a> <a id="2351" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2230" class="Function">∨-identityˡ</a> <a id="2363" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="2365" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2062" class="Function">∨-identityʳ</a>
|
||
|
||
<a id="∧-zeroʳ"></a><a id="2378" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2378" class="Function">∧-zeroʳ</a> <a id="2386" class="Symbol">:</a> <a id="2388" href="Algebra.Definitions.html#2007" class="Function">RightZero</a> <a id="2398" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="2400" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a>
|
||
<a id="2404" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2378" class="Function">∧-zeroʳ</a> <a id="2412" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2414" class="Symbol">=</a> <a id="2416" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="2424" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2426" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="2428" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="2439" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="2442" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="2450" class="Symbol">(</a><a id="2451" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="2465" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a><a id="2466" class="Symbol">)</a> <a id="2468" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="2472" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2474" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="2477" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2480" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="2482" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="2484" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2487" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="2490" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="2498" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2500" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2502" class="Symbol">(</a><a id="2503" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="2505" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a><a id="2506" class="Symbol">)</a> <a id="2508" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="2512" class="Symbol">(</a><a id="2513" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2515" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="2517" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a><a id="2518" class="Symbol">)</a> <a id="2520" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="2522" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="2524" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2527" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2531" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="2539" class="Symbol">(</a><a id="2540" href="Algebra.Lattice.Properties.Lattice.html#893" class="Function">∧-idem</a> <a id="2547" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a><a id="2548" class="Symbol">)</a> <a id="2550" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2554" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2562" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="2564" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="2566" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2569" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2573" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="2587" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2412" class="Bound">x</a> <a id="2589" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2593" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="2608" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="∧-zeroˡ"></a><a id="2611" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2611" class="Function">∧-zeroˡ</a> <a id="2619" class="Symbol">:</a> <a id="2621" href="Algebra.Definitions.html#1942" class="Function">LeftZero</a> <a id="2630" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="2632" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a>
|
||
<a id="2636" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2611" class="Function">∧-zeroˡ</a> <a id="2644" class="Symbol">=</a> <a id="2646" href="Algebra.Consequences.Setoid.html#4953" class="Function">comm∧zeʳ⇒zeˡ</a> <a id="2659" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="2666" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2378" class="Function">∧-zeroʳ</a>
|
||
|
||
<a id="∧-zero"></a><a id="2675" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2675" class="Function">∧-zero</a> <a id="2682" class="Symbol">:</a> <a id="2684" href="Algebra.Definitions.html#2074" class="Function">Zero</a> <a id="2689" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="2691" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a>
|
||
<a id="2695" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2675" class="Function">∧-zero</a> <a id="2702" class="Symbol">=</a> <a id="2704" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2611" class="Function">∧-zeroˡ</a> <a id="2712" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="2714" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2378" class="Function">∧-zeroʳ</a>
|
||
|
||
<a id="∨-zeroʳ"></a><a id="2723" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2723" class="Function">∨-zeroʳ</a> <a id="2731" class="Symbol">:</a> <a id="2733" class="Symbol">∀</a> <a id="2735" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2735" class="Bound">x</a> <a id="2737" class="Symbol">→</a> <a id="2739" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2735" class="Bound">x</a> <a id="2741" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="2743" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="2745" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="2747" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a>
|
||
<a id="2749" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2723" class="Function">∨-zeroʳ</a> <a id="2757" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2759" class="Symbol">=</a> <a id="2761" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="2769" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2771" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="2773" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="2784" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="2787" href="Algebra.Lattice.Structures.html#4639" class="Function">∨-congˡ</a> <a id="2795" class="Symbol">(</a><a id="2796" href="Algebra.Lattice.Structures.html#5812" class="Function">∨-complementʳ</a> <a id="2810" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a><a id="2811" class="Symbol">)</a> <a id="2813" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="2817" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2819" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="2822" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2825" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="2827" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="2829" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2832" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="2835" href="Algebra.Lattice.Structures.html#4147" class="Function">∨-assoc</a> <a id="2843" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2845" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2847" class="Symbol">(</a><a id="2848" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="2850" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a><a id="2851" class="Symbol">)</a> <a id="2853" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="2857" class="Symbol">(</a><a id="2858" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2860" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="2862" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a><a id="2863" class="Symbol">)</a> <a id="2865" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="2867" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="2869" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2872" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2875" href="Algebra.Lattice.Structures.html#4700" class="Function">∨-congʳ</a> <a id="2883" class="Symbol">(</a><a id="2884" href="Algebra.Lattice.Properties.Lattice.html#1877" class="Function">∨-idem</a> <a id="2891" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a><a id="2892" class="Symbol">)</a> <a id="2894" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2898" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2906" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="2908" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="2910" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2913" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="2916" href="Algebra.Lattice.Structures.html#5812" class="Function">∨-complementʳ</a> <a id="2930" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2757" class="Bound">x</a> <a id="2932" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="2936" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="2951" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="∨-zeroˡ"></a><a id="2954" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2954" class="Function">∨-zeroˡ</a> <a id="2962" class="Symbol">:</a> <a id="2964" href="Algebra.Definitions.html#1942" class="Function">LeftZero</a> <a id="2973" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="2975" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a>
|
||
<a id="2979" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2954" class="Function">∨-zeroˡ</a> <a id="2987" class="Symbol">=</a> <a id="2989" href="Algebra.Consequences.Setoid.html#4953" class="Function">comm∧zeʳ⇒zeˡ</a> <a id="3002" href="Algebra.Lattice.Structures.html#4113" class="Function">∨-comm</a> <a id="3009" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2723" class="Function">∨-zeroʳ</a>
|
||
|
||
<a id="∨-zero"></a><a id="3018" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3018" class="Function">∨-zero</a> <a id="3025" class="Symbol">:</a> <a id="3027" href="Algebra.Definitions.html#2074" class="Function">Zero</a> <a id="3032" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="3034" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a>
|
||
<a id="3038" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3018" class="Function">∨-zero</a> <a id="3045" class="Symbol">=</a> <a id="3047" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2954" class="Function">∨-zeroˡ</a> <a id="3055" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="3057" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2723" class="Function">∨-zeroʳ</a>
|
||
|
||
<a id="∨-⊥-isMonoid"></a><a id="3066" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3066" class="Function">∨-⊥-isMonoid</a> <a id="3079" class="Symbol">:</a> <a id="3081" href="Algebra.Structures.html#3974" class="Record">IsMonoid</a> <a id="3090" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a> <a id="3094" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a>
|
||
<a id="3096" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3066" class="Function">∨-⊥-isMonoid</a> <a id="3109" class="Symbol">=</a> <a id="3111" class="Keyword">record</a>
|
||
<a id="3120" class="Symbol">{</a> <a id="3122" href="Algebra.Structures.html#4035" class="Field">isSemigroup</a> <a id="3134" class="Symbol">=</a> <a id="3136" href="Algebra.Lattice.Properties.Lattice.html#2115" class="Function">∨-isSemigroup</a>
|
||
<a id="3152" class="Symbol">;</a> <a id="3154" href="Algebra.Structures.html#4067" class="Field">identity</a> <a id="3166" class="Symbol">=</a> <a id="3168" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2310" class="Function">∨-identity</a>
|
||
<a id="3181" class="Symbol">}</a>
|
||
|
||
<a id="∧-⊤-isMonoid"></a><a id="3184" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3184" class="Function">∧-⊤-isMonoid</a> <a id="3197" class="Symbol">:</a> <a id="3199" href="Algebra.Structures.html#3974" class="Record">IsMonoid</a> <a id="3208" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="3212" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a>
|
||
<a id="3214" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3184" class="Function">∧-⊤-isMonoid</a> <a id="3227" class="Symbol">=</a> <a id="3229" class="Keyword">record</a>
|
||
<a id="3238" class="Symbol">{</a> <a id="3240" href="Algebra.Structures.html#4035" class="Field">isSemigroup</a> <a id="3252" class="Symbol">=</a> <a id="3254" href="Algebra.Lattice.Properties.Lattice.html#1156" class="Function">∧-isSemigroup</a>
|
||
<a id="3270" class="Symbol">;</a> <a id="3272" href="Algebra.Structures.html#4067" class="Field">identity</a> <a id="3284" class="Symbol">=</a> <a id="3286" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1994" class="Function">∧-identity</a>
|
||
<a id="3299" class="Symbol">}</a>
|
||
|
||
<a id="∨-⊥-isCommutativeMonoid"></a><a id="3302" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3302" class="Function">∨-⊥-isCommutativeMonoid</a> <a id="3326" class="Symbol">:</a> <a id="3328" href="Algebra.Structures.html#4384" class="Record">IsCommutativeMonoid</a> <a id="3348" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a> <a id="3352" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a>
|
||
<a id="3354" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3302" class="Function">∨-⊥-isCommutativeMonoid</a> <a id="3378" class="Symbol">=</a> <a id="3380" class="Keyword">record</a>
|
||
<a id="3389" class="Symbol">{</a> <a id="3391" href="Algebra.Structures.html#4456" class="Field">isMonoid</a> <a id="3400" class="Symbol">=</a> <a id="3402" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3066" class="Function">∨-⊥-isMonoid</a>
|
||
<a id="3417" class="Symbol">;</a> <a id="3419" href="Algebra.Structures.html#4484" class="Field">comm</a> <a id="3428" class="Symbol">=</a> <a id="3430" href="Algebra.Lattice.Structures.html#4113" class="Function">∨-comm</a>
|
||
<a id="3439" class="Symbol">}</a>
|
||
|
||
<a id="∧-⊤-isCommutativeMonoid"></a><a id="3442" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3442" class="Function">∧-⊤-isCommutativeMonoid</a> <a id="3466" class="Symbol">:</a> <a id="3468" href="Algebra.Structures.html#4384" class="Record">IsCommutativeMonoid</a> <a id="3488" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="3492" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a>
|
||
<a id="3494" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3442" class="Function">∧-⊤-isCommutativeMonoid</a> <a id="3518" class="Symbol">=</a> <a id="3520" class="Keyword">record</a>
|
||
<a id="3529" class="Symbol">{</a> <a id="3531" href="Algebra.Structures.html#4456" class="Field">isMonoid</a> <a id="3540" class="Symbol">=</a> <a id="3542" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3184" class="Function">∧-⊤-isMonoid</a>
|
||
<a id="3557" class="Symbol">;</a> <a id="3559" href="Algebra.Structures.html#4484" class="Field">comm</a> <a id="3568" class="Symbol">=</a> <a id="3570" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a>
|
||
<a id="3579" class="Symbol">}</a>
|
||
|
||
<a id="∨-∧-isSemiring"></a><a id="3582" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3582" class="Function">∨-∧-isSemiring</a> <a id="3597" class="Symbol">:</a> <a id="3599" href="Algebra.Structures.html#12731" class="Record">IsSemiring</a> <a id="3610" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a> <a id="3614" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="3618" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="3620" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a>
|
||
<a id="3622" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3582" class="Function">∨-∧-isSemiring</a> <a id="3637" class="Symbol">=</a> <a id="3639" class="Keyword">record</a>
|
||
<a id="3648" class="Symbol">{</a> <a id="3650" href="Algebra.Structures.html#12800" class="Field">isSemiringWithoutAnnihilatingZero</a> <a id="3684" class="Symbol">=</a> <a id="3686" class="Keyword">record</a>
|
||
<a id="3697" class="Symbol">{</a> <a id="3699" href="Algebra.Structures.html#11137" class="Field">+-isCommutativeMonoid</a> <a id="3721" class="Symbol">=</a> <a id="3723" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3302" class="Function">∨-⊥-isCommutativeMonoid</a>
|
||
<a id="3751" class="Symbol">;</a> <a id="3753" href="Algebra.Structures.html#11190" class="Field">*-cong</a> <a id="3760" class="Symbol">=</a> <a id="3762" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a>
|
||
<a id="3773" class="Symbol">;</a> <a id="3775" href="Algebra.Structures.html#11231" class="Field">*-assoc</a> <a id="3783" class="Symbol">=</a> <a id="3785" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a>
|
||
<a id="3797" class="Symbol">;</a> <a id="3799" href="Algebra.Structures.html#11273" class="Field">*-identity</a> <a id="3810" class="Symbol">=</a> <a id="3812" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1994" class="Function">∧-identity</a>
|
||
<a id="3827" class="Symbol">;</a> <a id="3829" href="Algebra.Structures.html#11315" class="Field">distrib</a> <a id="3837" class="Symbol">=</a> <a id="3839" href="Algebra.Lattice.Structures.html#4906" class="Function">∧-distrib-∨</a>
|
||
<a id="3855" class="Symbol">}</a>
|
||
<a id="3859" class="Symbol">;</a> <a id="3861" href="Algebra.Structures.html#12890" class="Field">zero</a> <a id="3866" class="Symbol">=</a> <a id="3868" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2675" class="Function">∧-zero</a>
|
||
<a id="3877" class="Symbol">}</a>
|
||
|
||
<a id="∧-∨-isSemiring"></a><a id="3880" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3880" class="Function">∧-∨-isSemiring</a> <a id="3895" class="Symbol">:</a> <a id="3897" href="Algebra.Structures.html#12731" class="Record">IsSemiring</a> <a id="3908" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="3912" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a> <a id="3916" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="3918" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a>
|
||
<a id="3920" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3880" class="Function">∧-∨-isSemiring</a> <a id="3935" class="Symbol">=</a> <a id="3937" class="Keyword">record</a>
|
||
<a id="3946" class="Symbol">{</a> <a id="3948" href="Algebra.Structures.html#12800" class="Field">isSemiringWithoutAnnihilatingZero</a> <a id="3982" class="Symbol">=</a> <a id="3984" class="Keyword">record</a>
|
||
<a id="3995" class="Symbol">{</a> <a id="3997" href="Algebra.Structures.html#11137" class="Field">+-isCommutativeMonoid</a> <a id="4019" class="Symbol">=</a> <a id="4021" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3442" class="Function">∧-⊤-isCommutativeMonoid</a>
|
||
<a id="4049" class="Symbol">;</a> <a id="4051" href="Algebra.Structures.html#11190" class="Field">*-cong</a> <a id="4058" class="Symbol">=</a> <a id="4060" href="Algebra.Lattice.Structures.html#4181" class="Function">∨-cong</a>
|
||
<a id="4071" class="Symbol">;</a> <a id="4073" href="Algebra.Structures.html#11231" class="Field">*-assoc</a> <a id="4081" class="Symbol">=</a> <a id="4083" href="Algebra.Lattice.Structures.html#4147" class="Function">∨-assoc</a>
|
||
<a id="4095" class="Symbol">;</a> <a id="4097" href="Algebra.Structures.html#11273" class="Field">*-identity</a> <a id="4108" class="Symbol">=</a> <a id="4110" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2310" class="Function">∨-identity</a>
|
||
<a id="4125" class="Symbol">;</a> <a id="4127" href="Algebra.Structures.html#11315" class="Field">distrib</a> <a id="4135" class="Symbol">=</a> <a id="4137" href="Algebra.Lattice.Structures.html#4868" class="Function">∨-distrib-∧</a>
|
||
<a id="4153" class="Symbol">}</a>
|
||
<a id="4157" class="Symbol">;</a> <a id="4159" href="Algebra.Structures.html#12890" class="Field">zero</a> <a id="4164" class="Symbol">=</a> <a id="4166" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3018" class="Function">∨-zero</a>
|
||
<a id="4175" class="Symbol">}</a>
|
||
|
||
<a id="∨-∧-isCommutativeSemiring"></a><a id="4178" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4178" class="Function">∨-∧-isCommutativeSemiring</a> <a id="4204" class="Symbol">:</a> <a id="4206" href="Algebra.Structures.html#13418" class="Record">IsCommutativeSemiring</a> <a id="4228" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a> <a id="4232" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="4236" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="4238" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a>
|
||
<a id="4240" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4178" class="Function">∨-∧-isCommutativeSemiring</a> <a id="4266" class="Symbol">=</a> <a id="4268" class="Keyword">record</a>
|
||
<a id="4277" class="Symbol">{</a> <a id="4279" href="Algebra.Structures.html#13498" class="Field">isSemiring</a> <a id="4290" class="Symbol">=</a> <a id="4292" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3582" class="Function">∨-∧-isSemiring</a>
|
||
<a id="4309" class="Symbol">;</a> <a id="4311" href="Algebra.Structures.html#13536" class="Field">*-comm</a> <a id="4318" class="Symbol">=</a> <a id="4320" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a>
|
||
<a id="4329" class="Symbol">}</a>
|
||
|
||
<a id="∧-∨-isCommutativeSemiring"></a><a id="4332" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4332" class="Function">∧-∨-isCommutativeSemiring</a> <a id="4358" class="Symbol">:</a> <a id="4360" href="Algebra.Structures.html#13418" class="Record">IsCommutativeSemiring</a> <a id="4382" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="4386" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">_∨_</a> <a id="4390" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="4392" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a>
|
||
<a id="4394" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4332" class="Function">∧-∨-isCommutativeSemiring</a> <a id="4420" class="Symbol">=</a> <a id="4422" class="Keyword">record</a>
|
||
<a id="4431" class="Symbol">{</a> <a id="4433" href="Algebra.Structures.html#13498" class="Field">isSemiring</a> <a id="4444" class="Symbol">=</a> <a id="4446" href="Algebra.Lattice.Properties.BooleanAlgebra.html#3880" class="Function">∧-∨-isSemiring</a>
|
||
<a id="4463" class="Symbol">;</a> <a id="4465" href="Algebra.Structures.html#13536" class="Field">*-comm</a> <a id="4472" class="Symbol">=</a> <a id="4474" href="Algebra.Lattice.Structures.html#4113" class="Function">∨-comm</a>
|
||
<a id="4483" class="Symbol">}</a>
|
||
|
||
<a id="∨-∧-commutativeSemiring"></a><a id="4486" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4486" class="Function">∨-∧-commutativeSemiring</a> <a id="4510" class="Symbol">:</a> <a id="4512" href="Algebra.Bundles.html#17471" class="Record">CommutativeSemiring</a> <a id="4532" class="Symbol">_</a> <a id="4534" class="Symbol">_</a>
|
||
<a id="4536" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4486" class="Function">∨-∧-commutativeSemiring</a> <a id="4560" class="Symbol">=</a> <a id="4562" class="Keyword">record</a>
|
||
<a id="4571" class="Symbol">{</a> <a id="4573" href="Algebra.Bundles.html#17806" class="Field">isCommutativeSemiring</a> <a id="4595" class="Symbol">=</a> <a id="4597" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4178" class="Function">∨-∧-isCommutativeSemiring</a>
|
||
<a id="4625" class="Symbol">}</a>
|
||
|
||
<a id="∧-∨-commutativeSemiring"></a><a id="4628" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4628" class="Function">∧-∨-commutativeSemiring</a> <a id="4652" class="Symbol">:</a> <a id="4654" href="Algebra.Bundles.html#17471" class="Record">CommutativeSemiring</a> <a id="4674" class="Symbol">_</a> <a id="4676" class="Symbol">_</a>
|
||
<a id="4678" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4628" class="Function">∧-∨-commutativeSemiring</a> <a id="4702" class="Symbol">=</a> <a id="4704" class="Keyword">record</a>
|
||
<a id="4713" class="Symbol">{</a> <a id="4715" href="Algebra.Bundles.html#17806" class="Field">isCommutativeSemiring</a> <a id="4737" class="Symbol">=</a> <a id="4739" href="Algebra.Lattice.Properties.BooleanAlgebra.html#4332" class="Function">∧-∨-isCommutativeSemiring</a>
|
||
<a id="4767" class="Symbol">}</a>
|
||
|
||
<a id="4770" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="4843" class="Comment">-- Some other properties</a>
|
||
|
||
<a id="4869" class="Comment">-- I took the statement of this lemma (called Uniqueness of</a>
|
||
<a id="4929" class="Comment">-- Complements) from some course notes, "Boolean Algebra", written</a>
|
||
<a id="4996" class="Comment">-- by Gert Smolka.</a>
|
||
|
||
<a id="5016" class="Keyword">private</a>
|
||
<a id="lemma"></a><a id="5026" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5026" class="Function">lemma</a> <a id="5032" class="Symbol">:</a> <a id="5034" class="Symbol">∀</a> <a id="5036" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5036" class="Bound">x</a> <a id="5038" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5038" class="Bound">y</a> <a id="5040" class="Symbol">→</a> <a id="5042" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5036" class="Bound">x</a> <a id="5044" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5046" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5038" class="Bound">y</a> <a id="5048" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="5050" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="5052" class="Symbol">→</a> <a id="5054" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5036" class="Bound">x</a> <a id="5056" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5058" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5038" class="Bound">y</a> <a id="5060" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="5062" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="5064" class="Symbol">→</a> <a id="5066" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5068" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5036" class="Bound">x</a> <a id="5070" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="5072" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5038" class="Bound">y</a>
|
||
<a id="5076" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5026" class="Function">lemma</a> <a id="5082" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5084" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5084" class="Bound">y</a> <a id="5086" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5086" class="Bound">x∧y=⊥</a> <a id="5092" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5092" class="Bound">x∨y=⊤</a> <a id="5098" class="Symbol">=</a> <a id="5100" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="5110" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5112" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5129" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="5132" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1746" class="Function">∧-identityʳ</a> <a id="5144" class="Symbol">_</a> <a id="5146" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="5152" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5154" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5156" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5158" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="5171" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="5174" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="5182" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5092" class="Bound">x∨y=⊤</a> <a id="5188" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="5194" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5196" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5198" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5200" class="Symbol">(</a><a id="5201" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5203" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5205" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5084" class="Bound">y</a><a id="5206" class="Symbol">)</a> <a id="5213" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5217" href="Algebra.Lattice.Structures.html#5126" class="Function">∧-distribˡ-∨</a> <a id="5230" class="Symbol">_</a> <a id="5232" class="Symbol">_</a> <a id="5234" class="Symbol">_</a> <a id="5236" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="5242" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5244" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5246" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5248" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5250" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5252" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5254" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5256" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5258" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5084" class="Bound">y</a> <a id="5261" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5265" href="Algebra.Lattice.Structures.html#4700" class="Function">∨-congʳ</a> <a id="5273" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="5275" href="Algebra.Lattice.Structures.html#5887" class="Function">∧-complementˡ</a> <a id="5289" class="Symbol">_</a> <a id="5291" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="5297" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="5299" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5301" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5303" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5305" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5307" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5084" class="Bound">y</a> <a id="5316" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="5319" href="Algebra.Lattice.Structures.html#4700" class="Function">∨-congʳ</a> <a id="5327" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5086" class="Bound">x∧y=⊥</a> <a id="5333" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="5339" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5341" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5343" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5084" class="Bound">y</a> <a id="5345" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5347" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5349" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5351" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5353" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5084" class="Bound">y</a> <a id="5358" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="5361" href="Algebra.Lattice.Structures.html#5200" class="Function">∧-distribʳ-∨</a> <a id="5374" class="Symbol">_</a> <a id="5376" class="Symbol">_</a> <a id="5378" class="Symbol">_</a> <a id="5380" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="5386" class="Symbol">(</a><a id="5387" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a> <a id="5389" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5391" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5393" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5082" class="Bound">x</a><a id="5394" class="Symbol">)</a> <a id="5396" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5398" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5084" class="Bound">y</a> <a id="5405" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5409" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="5417" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="5419" href="Algebra.Lattice.Structures.html#5812" class="Function">∨-complementʳ</a> <a id="5433" class="Symbol">_</a> <a id="5435" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="5441" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="5443" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5445" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5084" class="Bound">y</a> <a id="5460" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5464" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1914" class="Function">∧-identityˡ</a> <a id="5476" class="Symbol">_</a> <a id="5478" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="5484" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5084" class="Bound">y</a> <a id="5503" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="⊥≉⊤"></a><a id="5506" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5506" class="Function">⊥≉⊤</a> <a id="5510" class="Symbol">:</a> <a id="5512" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5514" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="5516" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="5518" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a>
|
||
<a id="5520" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5506" class="Function">⊥≉⊤</a> <a id="5524" class="Symbol">=</a> <a id="5526" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5026" class="Function">lemma</a> <a id="5532" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="5534" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="5536" class="Symbol">(</a><a id="5537" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1746" class="Function">∧-identityʳ</a> <a id="5549" class="Symbol">_)</a> <a id="5552" class="Symbol">(</a><a id="5553" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2723" class="Function">∨-zeroʳ</a> <a id="5561" class="Symbol">_)</a>
|
||
|
||
<a id="⊤≉⊥"></a><a id="5565" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5565" class="Function">⊤≉⊥</a> <a id="5569" class="Symbol">:</a> <a id="5571" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5573" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="5575" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="5577" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a>
|
||
<a id="5579" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5565" class="Function">⊤≉⊥</a> <a id="5583" class="Symbol">=</a> <a id="5585" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5026" class="Function">lemma</a> <a id="5591" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="5593" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="5595" class="Symbol">(</a><a id="5596" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2378" class="Function">∧-zeroʳ</a> <a id="5604" class="Symbol">_)</a> <a id="5607" class="Symbol">(</a><a id="5608" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2062" class="Function">∨-identityʳ</a> <a id="5620" class="Symbol">_)</a>
|
||
|
||
<a id="¬-involutive"></a><a id="5624" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5624" class="Function">¬-involutive</a> <a id="5637" class="Symbol">:</a> <a id="5639" href="Algebra.Definitions.html#4153" class="Function">Involutive</a> <a id="5650" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬_</a>
|
||
<a id="5653" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5624" class="Function">¬-involutive</a> <a id="5666" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5666" class="Bound">x</a> <a id="5668" class="Symbol">=</a> <a id="5670" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5026" class="Function">lemma</a> <a id="5676" class="Symbol">(</a><a id="5677" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5679" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5666" class="Bound">x</a><a id="5680" class="Symbol">)</a> <a id="5682" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5666" class="Bound">x</a> <a id="5684" class="Symbol">(</a><a id="5685" href="Algebra.Lattice.Structures.html#5887" class="Function">∧-complementˡ</a> <a id="5699" class="Symbol">_)</a> <a id="5702" class="Symbol">(</a><a id="5703" href="Algebra.Lattice.Structures.html#5738" class="Function">∨-complementˡ</a> <a id="5717" class="Symbol">_)</a>
|
||
|
||
<a id="deMorgan₁"></a><a id="5721" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="5731" class="Symbol">:</a> <a id="5733" class="Symbol">∀</a> <a id="5735" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5735" class="Bound">x</a> <a id="5737" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5737" class="Bound">y</a> <a id="5739" class="Symbol">→</a> <a id="5741" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5743" class="Symbol">(</a><a id="5744" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5735" class="Bound">x</a> <a id="5746" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5748" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5737" class="Bound">y</a><a id="5749" class="Symbol">)</a> <a id="5751" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="5753" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5755" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5735" class="Bound">x</a> <a id="5757" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5759" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5761" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5737" class="Bound">y</a>
|
||
<a id="5763" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="5773" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5775" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="5777" class="Symbol">=</a> <a id="5779" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5026" class="Function">lemma</a> <a id="5785" class="Symbol">(</a><a id="5786" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5788" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5790" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="5791" class="Symbol">)</a> <a id="5793" class="Symbol">(</a><a id="5794" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5796" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5798" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5800" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5802" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="5803" class="Symbol">)</a> <a id="5805" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5825" class="Function">lem₁</a> <a id="5810" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6600" class="Function">lem₂</a>
|
||
<a id="5817" class="Keyword">where</a>
|
||
<a id="5825" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5825" class="Function">lem₁</a> <a id="5830" class="Symbol">=</a> <a id="5832" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="5842" class="Symbol">(</a><a id="5843" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5845" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5847" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="5848" class="Symbol">)</a> <a id="5850" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5852" class="Symbol">(</a><a id="5853" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5855" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5857" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5859" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5861" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="5862" class="Symbol">)</a> <a id="5873" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5876" href="Algebra.Lattice.Structures.html#5126" class="Function">∧-distribˡ-∨</a> <a id="5889" class="Symbol">_</a> <a id="5891" class="Symbol">_</a> <a id="5893" class="Symbol">_</a> <a id="5895" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="5901" class="Symbol">(</a><a id="5902" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5904" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5906" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="5907" class="Symbol">)</a> <a id="5909" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5911" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5913" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5915" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5917" class="Symbol">(</a><a id="5918" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5920" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5922" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="5923" class="Symbol">)</a> <a id="5925" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5927" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5929" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="5932" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="5935" href="Algebra.Lattice.Structures.html#4700" class="Function">∨-congʳ</a> <a id="5943" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="5945" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="5953" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="5955" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="5962" class="Symbol">_</a> <a id="5964" class="Symbol">_</a> <a id="5966" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="5972" class="Symbol">(</a><a id="5973" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="5975" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5977" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a><a id="5978" class="Symbol">)</a> <a id="5980" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5982" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="5984" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5986" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="5988" class="Symbol">(</a><a id="5989" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="5991" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5993" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="5994" class="Symbol">)</a> <a id="5996" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="5998" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6000" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6003" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6006" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="6014" class="Symbol">_</a> <a id="6016" class="Symbol">_</a> <a id="6018" class="Symbol">_</a> <a id="6020" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="6022" href="Algebra.Lattice.Structures.html#4181" class="Function">∨-cong</a> <a id="6029" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="6031" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="6039" class="Symbol">_</a> <a id="6041" class="Symbol">_</a> <a id="6043" class="Symbol">_</a> <a id="6045" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6051" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6053" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6055" class="Symbol">(</a><a id="6056" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6058" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6060" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6062" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a><a id="6063" class="Symbol">)</a> <a id="6065" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6067" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6069" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6071" class="Symbol">(</a><a id="6072" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6074" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6076" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6078" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="6079" class="Symbol">)</a> <a id="6082" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6085" class="Symbol">(</a><a id="6086" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="6094" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="6096" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="6110" class="Symbol">_)</a> <a id="6113" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="6115" href="Algebra.Lattice.Structures.html#4181" class="Function">∨-cong</a> <a id="6122" href="Function.Base.html#4322" class="Function Operator">⟩</a>
|
||
<a id="6162" class="Symbol">(</a><a id="6163" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="6171" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="6173" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="6187" class="Symbol">_)</a> <a id="6190" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6196" class="Symbol">(</a><a id="6197" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6199" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6201" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a><a id="6202" class="Symbol">)</a> <a id="6204" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6206" class="Symbol">(</a><a id="6207" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6209" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6211" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a><a id="6212" class="Symbol">)</a> <a id="6227" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6230" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2378" class="Function">∧-zeroʳ</a> <a id="6238" class="Symbol">_</a> <a id="6240" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="6242" href="Algebra.Lattice.Structures.html#4181" class="Function">∨-cong</a> <a id="6249" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="6251" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2378" class="Function">∧-zeroʳ</a> <a id="6259" class="Symbol">_</a> <a id="6261" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6267" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="6269" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6271" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="6298" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6301" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2062" class="Function">∨-identityʳ</a> <a id="6313" class="Symbol">_</a> <a id="6315" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6321" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="6352" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="6357" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6357" class="Function">lem₃</a> <a id="6362" class="Symbol">=</a> <a id="6364" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="6374" class="Symbol">(</a><a id="6375" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6377" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6379" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="6380" class="Symbol">)</a> <a id="6382" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6384" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6386" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6397" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6400" href="Algebra.Lattice.Structures.html#5052" class="Function">∨-distribʳ-∧</a> <a id="6413" class="Symbol">_</a> <a id="6415" class="Symbol">_</a> <a id="6417" class="Symbol">_</a> <a id="6419" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6425" class="Symbol">(</a><a id="6426" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6428" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6430" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6432" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a><a id="6433" class="Symbol">)</a> <a id="6435" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6437" class="Symbol">(</a><a id="6438" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6440" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6442" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6444" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a><a id="6445" class="Symbol">)</a> <a id="6448" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6451" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="6459" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="6461" href="Algebra.Lattice.Structures.html#5812" class="Function">∨-complementʳ</a> <a id="6475" class="Symbol">_</a> <a id="6477" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6483" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="6485" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6487" class="Symbol">(</a><a id="6488" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6490" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6492" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6494" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a><a id="6495" class="Symbol">)</a> <a id="6506" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6509" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1914" class="Function">∧-identityˡ</a> <a id="6521" class="Symbol">_</a> <a id="6523" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6529" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6531" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6533" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6535" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6552" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6555" href="Algebra.Lattice.Structures.html#4113" class="Function">∨-comm</a> <a id="6562" class="Symbol">_</a> <a id="6564" class="Symbol">_</a> <a id="6566" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6572" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6574" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6576" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6578" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6595" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="6600" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6600" class="Function">lem₂</a> <a id="6605" class="Symbol">=</a> <a id="6607" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="6617" class="Symbol">(</a><a id="6618" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6620" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6622" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="6623" class="Symbol">)</a> <a id="6625" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6627" class="Symbol">(</a><a id="6628" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6630" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6632" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6634" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6636" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="6637" class="Symbol">)</a> <a id="6640" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="6643" href="Algebra.Lattice.Structures.html#4147" class="Function">∨-assoc</a> <a id="6651" class="Symbol">_</a> <a id="6653" class="Symbol">_</a> <a id="6655" class="Symbol">_</a> <a id="6657" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="6663" class="Symbol">((</a><a id="6665" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6667" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6669" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="6670" class="Symbol">)</a> <a id="6672" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6674" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6676" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a><a id="6677" class="Symbol">)</a> <a id="6679" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6681" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6683" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6686" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6689" href="Algebra.Lattice.Structures.html#4700" class="Function">∨-congʳ</a> <a id="6697" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6357" class="Function">lem₃</a> <a id="6702" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6708" class="Symbol">(</a><a id="6709" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6711" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6713" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6715" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="6716" class="Symbol">)</a> <a id="6718" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6720" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6722" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6731" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6734" href="Algebra.Lattice.Structures.html#4147" class="Function">∨-assoc</a> <a id="6742" class="Symbol">_</a> <a id="6744" class="Symbol">_</a> <a id="6746" class="Symbol">_</a> <a id="6748" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6754" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6756" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6758" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6760" class="Symbol">(</a><a id="6761" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a> <a id="6763" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6765" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6767" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5775" class="Bound">y</a><a id="6768" class="Symbol">)</a> <a id="6777" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6780" href="Algebra.Lattice.Structures.html#4639" class="Function">∨-congˡ</a> <a id="6788" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="6790" href="Algebra.Lattice.Structures.html#5812" class="Function">∨-complementʳ</a> <a id="6804" class="Symbol">_</a> <a id="6806" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6812" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6814" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5773" class="Bound">x</a> <a id="6816" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6818" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="6835" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="6838" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2723" class="Function">∨-zeroʳ</a> <a id="6846" class="Symbol">_</a> <a id="6848" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="6854" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="6877" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="deMorgan₂"></a><a id="6880" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6880" class="Function">deMorgan₂</a> <a id="6890" class="Symbol">:</a> <a id="6892" class="Symbol">∀</a> <a id="6894" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6894" class="Bound">x</a> <a id="6896" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6896" class="Bound">y</a> <a id="6898" class="Symbol">→</a> <a id="6900" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6902" class="Symbol">(</a><a id="6903" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6894" class="Bound">x</a> <a id="6905" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6907" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6896" class="Bound">y</a><a id="6908" class="Symbol">)</a> <a id="6910" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="6912" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6914" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6894" class="Bound">x</a> <a id="6916" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="6918" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6920" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6896" class="Bound">y</a>
|
||
<a id="6922" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6880" class="Function">deMorgan₂</a> <a id="6932" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6932" class="Bound">x</a> <a id="6934" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6934" class="Bound">y</a> <a id="6936" class="Symbol">=</a> <a id="6938" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="6946" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="6948" class="Symbol">(</a><a id="6949" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6932" class="Bound">x</a> <a id="6951" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="6953" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6934" class="Bound">y</a><a id="6954" class="Symbol">)</a> <a id="6965" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="6968" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="6975" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="6977" class="Symbol">((</a><a id="6979" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5624" class="Function">¬-involutive</a> <a id="6992" class="Symbol">_)</a> <a id="6995" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="6997" href="Algebra.Lattice.Structures.html#4181" class="Function">∨-cong</a> <a id="7004" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="7006" class="Symbol">(</a><a id="7007" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5624" class="Function">¬-involutive</a> <a id="7020" class="Symbol">_))</a> <a id="7024" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="7028" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7030" class="Symbol">(</a><a id="7031" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7033" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7035" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6932" class="Bound">x</a> <a id="7037" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="7039" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7041" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7043" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6934" class="Bound">y</a><a id="7044" class="Symbol">)</a> <a id="7047" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="7050" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="7057" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="7059" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="7069" class="Symbol">_</a> <a id="7071" class="Symbol">_</a> <a id="7073" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="7077" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7079" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7081" class="Symbol">(</a><a id="7082" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7084" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6932" class="Bound">x</a> <a id="7086" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7088" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7090" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6934" class="Bound">y</a><a id="7091" class="Symbol">)</a> <a id="7096" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="7099" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5624" class="Function">¬-involutive</a> <a id="7112" class="Symbol">_</a> <a id="7114" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="7118" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7120" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6932" class="Bound">x</a> <a id="7122" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7124" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7126" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6934" class="Bound">y</a> <a id="7137" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="7140" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="7213" class="Comment">-- (⊕, ∧, id, ⊥, ⊤) is a commutative ring</a>
|
||
|
||
<a id="7256" class="Comment">-- This construction is parameterised over the definition of xor.</a>
|
||
|
||
<a id="7323" class="Keyword">module</a> <a id="XorRing"></a><a id="7330" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7330" class="Module">XorRing</a>
|
||
<a id="7340" class="Symbol">(</a><a id="7341" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7341" class="Bound">xor</a> <a id="7345" class="Symbol">:</a> <a id="7347" href="Algebra.Core.html#527" class="Function">Op₂</a> <a id="7351" href="Algebra.Lattice.Bundles.html#5670" class="Field">Carrier</a><a id="7358" class="Symbol">)</a>
|
||
<a id="7362" class="Symbol">(</a><a id="7363" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="7369" class="Symbol">:</a> <a id="7371" class="Symbol">∀</a> <a id="7373" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7373" class="Bound">x</a> <a id="7375" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7375" class="Bound">y</a> <a id="7377" class="Symbol">→</a> <a id="7379" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7341" class="Bound">xor</a> <a id="7383" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7373" class="Bound">x</a> <a id="7385" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7375" class="Bound">y</a> <a id="7387" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="7389" class="Symbol">(</a><a id="7390" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7373" class="Bound">x</a> <a id="7392" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="7394" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7375" class="Bound">y</a><a id="7395" class="Symbol">)</a> <a id="7397" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7399" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7401" class="Symbol">(</a><a id="7402" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7373" class="Bound">x</a> <a id="7404" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7406" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7375" class="Bound">y</a><a id="7407" class="Symbol">))</a>
|
||
<a id="7412" class="Keyword">where</a>
|
||
|
||
<a id="7421" class="Keyword">private</a>
|
||
<a id="7433" class="Keyword">infixl</a> <a id="7440" class="Number">6</a> <a id="7442" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
|
||
<a id="XorRing._⊕_"></a><a id="7451" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a> <a id="7455" class="Symbol">:</a> <a id="7457" href="Algebra.Core.html#527" class="Function">Op₂</a> <a id="7461" href="Algebra.Lattice.Bundles.html#5670" class="Field">Carrier</a>
|
||
<a id="7473" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a> <a id="7477" class="Symbol">=</a> <a id="7479" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7341" class="Bound">xor</a>
|
||
|
||
<a id="XorRing.helper"></a><a id="7488" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7488" class="Function">helper</a> <a id="7495" class="Symbol">:</a> <a id="7497" class="Symbol">∀</a> <a id="7499" class="Symbol">{</a><a id="7500" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7500" class="Bound">x</a> <a id="7502" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7502" class="Bound">y</a> <a id="7504" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7504" class="Bound">u</a> <a id="7506" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7506" class="Bound">v</a><a id="7507" class="Symbol">}</a> <a id="7509" class="Symbol">→</a> <a id="7511" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7500" class="Bound">x</a> <a id="7513" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="7515" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7502" class="Bound">y</a> <a id="7517" class="Symbol">→</a> <a id="7519" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7504" class="Bound">u</a> <a id="7521" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="7523" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7506" class="Bound">v</a> <a id="7525" class="Symbol">→</a> <a id="7527" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7500" class="Bound">x</a> <a id="7529" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7531" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7533" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7504" class="Bound">u</a> <a id="7535" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="7537" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7502" class="Bound">y</a> <a id="7539" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7541" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7543" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7506" class="Bound">v</a>
|
||
<a id="7549" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7488" class="Function">helper</a> <a id="7556" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7556" class="Bound">x≈y</a> <a id="7560" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7560" class="Bound">u≈v</a> <a id="7564" class="Symbol">=</a> <a id="7566" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7556" class="Bound">x≈y</a> <a id="7570" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="7572" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="7579" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="7581" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="7588" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7560" class="Bound">u≈v</a>
|
||
|
||
<a id="XorRing.⊕-cong"></a><a id="7595" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7595" class="Function">⊕-cong</a> <a id="7602" class="Symbol">:</a> <a id="7604" href="Algebra.Definitions.html#1302" class="Function">Congruent₂</a> <a id="7615" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="7621" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7595" class="Function">⊕-cong</a> <a id="7628" class="Symbol">{</a><a id="7629" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7629" class="Bound">x</a><a id="7630" class="Symbol">}</a> <a id="7632" class="Symbol">{</a><a id="7633" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7633" class="Bound">y</a><a id="7634" class="Symbol">}</a> <a id="7636" class="Symbol">{</a><a id="7637" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7637" class="Bound">u</a><a id="7638" class="Symbol">}</a> <a id="7640" class="Symbol">{</a><a id="7641" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7641" class="Bound">v</a><a id="7642" class="Symbol">}</a> <a id="7644" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7644" class="Bound">x≈y</a> <a id="7648" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7648" class="Bound">u≈v</a> <a id="7652" class="Symbol">=</a> <a id="7654" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="7664" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7629" class="Bound">x</a> <a id="7666" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="7668" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7637" class="Bound">u</a> <a id="7685" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="7689" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="7695" class="Symbol">_</a> <a id="7697" class="Symbol">_</a> <a id="7699" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="7705" class="Symbol">(</a><a id="7706" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7629" class="Bound">x</a> <a id="7708" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="7710" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7637" class="Bound">u</a><a id="7711" class="Symbol">)</a> <a id="7713" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7715" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7717" class="Symbol">(</a><a id="7718" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7629" class="Bound">x</a> <a id="7720" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7722" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7637" class="Bound">u</a><a id="7723" class="Symbol">)</a> <a id="7726" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="7730" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7488" class="Function">helper</a> <a id="7737" class="Symbol">(</a><a id="7738" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7644" class="Bound">x≈y</a> <a id="7742" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="7744" href="Algebra.Lattice.Structures.html#4181" class="Function">∨-cong</a> <a id="7751" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="7753" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7648" class="Bound">u≈v</a><a id="7756" class="Symbol">)</a>
|
||
<a id="7794" class="Symbol">(</a><a id="7795" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7644" class="Bound">x≈y</a> <a id="7799" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="7801" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="7808" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="7810" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7648" class="Bound">u≈v</a><a id="7813" class="Symbol">)</a> <a id="7815" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="7821" class="Symbol">(</a><a id="7822" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7633" class="Bound">y</a> <a id="7824" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="7826" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7641" class="Bound">v</a><a id="7827" class="Symbol">)</a> <a id="7829" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7831" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7833" class="Symbol">(</a><a id="7834" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7633" class="Bound">y</a> <a id="7836" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7838" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7641" class="Bound">v</a><a id="7839" class="Symbol">)</a> <a id="7842" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="7845" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="7851" class="Symbol">_</a> <a id="7853" class="Symbol">_</a> <a id="7855" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="7861" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7633" class="Bound">y</a> <a id="7863" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="7865" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7641" class="Bound">v</a> <a id="7882" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.⊕-comm"></a><a id="7887" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7887" class="Function">⊕-comm</a> <a id="7894" class="Symbol">:</a> <a id="7896" href="Algebra.Definitions.html#1635" class="Function">Commutative</a> <a id="7908" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="7914" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7887" class="Function">⊕-comm</a> <a id="7921" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7921" class="Bound">x</a> <a id="7923" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7923" class="Bound">y</a> <a id="7925" class="Symbol">=</a> <a id="7927" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="7937" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7921" class="Bound">x</a> <a id="7939" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="7941" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7923" class="Bound">y</a> <a id="7958" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="7962" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="7968" class="Symbol">_</a> <a id="7970" class="Symbol">_</a> <a id="7972" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="7978" class="Symbol">(</a><a id="7979" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7921" class="Bound">x</a> <a id="7981" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="7983" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7923" class="Bound">y</a><a id="7984" class="Symbol">)</a> <a id="7986" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7988" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="7990" class="Symbol">(</a><a id="7991" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7921" class="Bound">x</a> <a id="7993" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="7995" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7923" class="Bound">y</a><a id="7996" class="Symbol">)</a> <a id="7999" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8003" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7488" class="Function">helper</a> <a id="8010" class="Symbol">(</a><a id="8011" href="Algebra.Lattice.Structures.html#4113" class="Function">∨-comm</a> <a id="8018" class="Symbol">_</a> <a id="8020" class="Symbol">_)</a> <a id="8023" class="Symbol">(</a><a id="8024" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="8031" class="Symbol">_</a> <a id="8033" class="Symbol">_)</a> <a id="8036" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8042" class="Symbol">(</a><a id="8043" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7923" class="Bound">y</a> <a id="8045" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="8047" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7921" class="Bound">x</a><a id="8048" class="Symbol">)</a> <a id="8050" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8052" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8054" class="Symbol">(</a><a id="8055" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7923" class="Bound">y</a> <a id="8057" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8059" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7921" class="Bound">x</a><a id="8060" class="Symbol">)</a> <a id="8063" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="8066" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="8072" class="Symbol">_</a> <a id="8074" class="Symbol">_</a> <a id="8076" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="8082" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7923" class="Bound">y</a> <a id="8084" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="8086" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7921" class="Bound">x</a> <a id="8103" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.¬-distribˡ-⊕"></a><a id="8108" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8108" class="Function">¬-distribˡ-⊕</a> <a id="8121" class="Symbol">:</a> <a id="8123" class="Symbol">∀</a> <a id="8125" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8125" class="Bound">x</a> <a id="8127" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8127" class="Bound">y</a> <a id="8129" class="Symbol">→</a> <a id="8131" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8133" class="Symbol">(</a><a id="8134" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8125" class="Bound">x</a> <a id="8136" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="8138" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8127" class="Bound">y</a><a id="8139" class="Symbol">)</a> <a id="8141" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="8143" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8145" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8125" class="Bound">x</a> <a id="8147" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="8149" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8127" class="Bound">y</a>
|
||
<a id="8153" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8108" class="Function">¬-distribˡ-⊕</a> <a id="8166" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8168" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a> <a id="8170" class="Symbol">=</a> <a id="8172" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="8182" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8184" class="Symbol">(</a><a id="8185" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8187" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="8189" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8190" class="Symbol">)</a> <a id="8221" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8224" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="8231" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="8233" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="8239" class="Symbol">_</a> <a id="8241" class="Symbol">_</a> <a id="8243" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8249" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8251" class="Symbol">((</a><a id="8253" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8255" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="8257" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8258" class="Symbol">)</a> <a id="8260" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8262" class="Symbol">(</a><a id="8263" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8265" class="Symbol">(</a><a id="8266" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8268" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8270" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8271" class="Symbol">)))</a> <a id="8288" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8291" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="8298" class="Symbol">(</a><a id="8299" href="Algebra.Lattice.Structures.html#5200" class="Function">∧-distribʳ-∨</a> <a id="8312" class="Symbol">_</a> <a id="8314" class="Symbol">_</a> <a id="8316" class="Symbol">_)</a> <a id="8319" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8325" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8327" class="Symbol">((</a><a id="8329" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8331" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8333" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8335" class="Symbol">(</a><a id="8336" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8338" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8340" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8341" class="Symbol">))</a> <a id="8344" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="8346" class="Symbol">(</a><a id="8347" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a> <a id="8349" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8351" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8353" class="Symbol">(</a><a id="8354" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8356" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8358" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8359" class="Symbol">)))</a> <a id="8364" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8367" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="8374" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="8376" href="Algebra.Lattice.Structures.html#4639" class="Function">∨-congˡ</a> <a id="8384" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="8386" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="8394" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="8396" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="8403" class="Symbol">(</a><a id="8404" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="8411" class="Symbol">_</a> <a id="8413" class="Symbol">_)</a> <a id="8416" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8422" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8424" class="Symbol">((</a><a id="8426" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8428" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8430" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8432" class="Symbol">(</a><a id="8433" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8435" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8437" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8438" class="Symbol">))</a> <a id="8441" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="8443" class="Symbol">(</a><a id="8444" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a> <a id="8446" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8448" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8450" class="Symbol">(</a><a id="8451" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a> <a id="8453" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8455" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a><a id="8456" class="Symbol">)))</a> <a id="8461" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8464" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="8471" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="8473" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8848" class="Function">lem</a> <a id="8477" class="Symbol">_</a> <a id="8479" class="Symbol">_</a> <a id="8481" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="8483" href="Algebra.Lattice.Structures.html#4181" class="Function">∨-cong</a> <a id="8490" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="8492" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8848" class="Function">lem</a> <a id="8496" class="Symbol">_</a> <a id="8498" class="Symbol">_</a> <a id="8500" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8506" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8508" class="Symbol">((</a><a id="8510" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8512" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8514" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8516" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8517" class="Symbol">)</a> <a id="8519" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="8521" class="Symbol">(</a><a id="8522" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a> <a id="8524" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8526" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8528" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a><a id="8529" class="Symbol">))</a> <a id="8545" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8548" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6880" class="Function">deMorgan₂</a> <a id="8558" class="Symbol">_</a> <a id="8560" class="Symbol">_</a> <a id="8562" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8568" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8570" class="Symbol">(</a><a id="8571" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8573" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8575" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8577" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8578" class="Symbol">)</a> <a id="8580" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8582" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8584" class="Symbol">(</a><a id="8585" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a> <a id="8587" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8589" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8591" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a><a id="8592" class="Symbol">)</a> <a id="8607" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8610" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="8618" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="8620" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="8630" class="Symbol">_</a> <a id="8632" class="Symbol">_</a> <a id="8634" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8640" class="Symbol">(</a><a id="8641" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8643" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8645" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="8647" class="Symbol">(</a><a id="8648" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8650" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8652" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8653" class="Symbol">))</a> <a id="8656" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8658" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8660" class="Symbol">(</a><a id="8661" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a> <a id="8663" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8665" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8667" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a><a id="8668" class="Symbol">)</a> <a id="8679" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8682" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7488" class="Function">helper</a> <a id="8689" class="Symbol">(</a><a id="8690" href="Algebra.Lattice.Structures.html#4639" class="Function">∨-congˡ</a> <a id="8698" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="8700" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5624" class="Function">¬-involutive</a> <a id="8713" class="Symbol">_)</a> <a id="8716" class="Symbol">(</a><a id="8717" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="8724" class="Symbol">_</a> <a id="8726" class="Symbol">_)</a> <a id="8729" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8735" class="Symbol">(</a><a id="8736" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8738" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8740" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="8742" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8743" class="Symbol">)</a> <a id="8745" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8747" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8749" class="Symbol">(</a><a id="8750" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8752" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8754" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8756" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a><a id="8757" class="Symbol">)</a> <a id="8774" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="8777" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="8783" class="Symbol">_</a> <a id="8785" class="Symbol">_</a> <a id="8787" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="8793" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8795" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8166" class="Bound">x</a> <a id="8797" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="8799" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8168" class="Bound">y</a> <a id="8832" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
<a id="8838" class="Keyword">where</a>
|
||
<a id="8848" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8848" class="Function">lem</a> <a id="8852" class="Symbol">:</a> <a id="8854" class="Symbol">∀</a> <a id="8856" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8856" class="Bound">x</a> <a id="8858" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8858" class="Bound">y</a> <a id="8860" class="Symbol">→</a> <a id="8862" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8856" class="Bound">x</a> <a id="8864" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8866" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8868" class="Symbol">(</a><a id="8869" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8856" class="Bound">x</a> <a id="8871" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8873" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8858" class="Bound">y</a><a id="8874" class="Symbol">)</a> <a id="8876" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="8878" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8856" class="Bound">x</a> <a id="8880" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8882" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8884" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8858" class="Bound">y</a>
|
||
<a id="8890" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8848" class="Function">lem</a> <a id="8894" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a> <a id="8896" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8896" class="Bound">y</a> <a id="8898" class="Symbol">=</a> <a id="8900" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="8912" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a> <a id="8914" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8916" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8918" class="Symbol">(</a><a id="8919" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a> <a id="8921" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8923" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8896" class="Bound">y</a><a id="8924" class="Symbol">)</a> <a id="8935" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8938" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="8946" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="8948" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="8958" class="Symbol">_</a> <a id="8960" class="Symbol">_</a> <a id="8962" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="8970" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a> <a id="8972" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="8974" class="Symbol">(</a><a id="8975" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8977" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a> <a id="8979" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="8981" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="8983" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8896" class="Bound">y</a><a id="8984" class="Symbol">)</a> <a id="8993" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="8996" href="Algebra.Lattice.Structures.html#5126" class="Function">∧-distribˡ-∨</a> <a id="9009" class="Symbol">_</a> <a id="9011" class="Symbol">_</a> <a id="9013" class="Symbol">_</a> <a id="9015" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9023" class="Symbol">(</a><a id="9024" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a> <a id="9026" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="9028" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9030" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a><a id="9031" class="Symbol">)</a> <a id="9033" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="9035" class="Symbol">(</a><a id="9036" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a> <a id="9038" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="9040" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9042" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8896" class="Bound">y</a><a id="9043" class="Symbol">)</a> <a id="9046" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9049" href="Algebra.Lattice.Structures.html#4700" class="Function">∨-congʳ</a> <a id="9057" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="9059" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="9073" class="Symbol">_</a> <a id="9075" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9083" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="9085" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="9087" class="Symbol">(</a><a id="9088" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a> <a id="9090" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="9092" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9094" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8896" class="Bound">y</a><a id="9095" class="Symbol">)</a> <a id="9106" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9109" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2230" class="Function">∨-identityˡ</a> <a id="9121" class="Symbol">_</a> <a id="9123" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9131" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8894" class="Bound">x</a> <a id="9133" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="9135" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9137" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8896" class="Bound">y</a> <a id="9154" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.¬-distribʳ-⊕"></a><a id="9159" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9159" class="Function">¬-distribʳ-⊕</a> <a id="9172" class="Symbol">:</a> <a id="9174" class="Symbol">∀</a> <a id="9176" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9176" class="Bound">x</a> <a id="9178" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9178" class="Bound">y</a> <a id="9180" class="Symbol">→</a> <a id="9182" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9184" class="Symbol">(</a><a id="9185" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9176" class="Bound">x</a> <a id="9187" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9189" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9178" class="Bound">y</a><a id="9190" class="Symbol">)</a> <a id="9192" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="9194" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9176" class="Bound">x</a> <a id="9196" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9198" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9200" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9178" class="Bound">y</a>
|
||
<a id="9204" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9159" class="Function">¬-distribʳ-⊕</a> <a id="9217" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9217" class="Bound">x</a> <a id="9219" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9219" class="Bound">y</a> <a id="9221" class="Symbol">=</a> <a id="9223" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="9233" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9235" class="Symbol">(</a><a id="9236" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9217" class="Bound">x</a> <a id="9238" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9240" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9219" class="Bound">y</a><a id="9241" class="Symbol">)</a> <a id="9244" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9247" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="9254" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="9256" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7887" class="Function">⊕-comm</a> <a id="9263" class="Symbol">_</a> <a id="9265" class="Symbol">_</a> <a id="9267" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9273" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9275" class="Symbol">(</a><a id="9276" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9219" class="Bound">y</a> <a id="9278" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9280" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9217" class="Bound">x</a><a id="9281" class="Symbol">)</a> <a id="9284" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9287" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8108" class="Function">¬-distribˡ-⊕</a> <a id="9300" class="Symbol">_</a> <a id="9302" class="Symbol">_</a> <a id="9304" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9310" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9312" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9219" class="Bound">y</a> <a id="9314" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9316" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9217" class="Bound">x</a> <a id="9321" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9324" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7887" class="Function">⊕-comm</a> <a id="9331" class="Symbol">_</a> <a id="9333" class="Symbol">_</a> <a id="9335" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9341" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9217" class="Bound">x</a> <a id="9343" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9345" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9347" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9219" class="Bound">y</a> <a id="9352" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.⊕-annihilates-¬"></a><a id="9357" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9357" class="Function">⊕-annihilates-¬</a> <a id="9373" class="Symbol">:</a> <a id="9375" class="Symbol">∀</a> <a id="9377" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9377" class="Bound">x</a> <a id="9379" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9379" class="Bound">y</a> <a id="9381" class="Symbol">→</a> <a id="9383" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9377" class="Bound">x</a> <a id="9385" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9387" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9379" class="Bound">y</a> <a id="9389" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a> <a id="9391" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9393" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9377" class="Bound">x</a> <a id="9395" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9397" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9399" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9379" class="Bound">y</a>
|
||
<a id="9403" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9357" class="Function">⊕-annihilates-¬</a> <a id="9419" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9419" class="Bound">x</a> <a id="9421" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9421" class="Bound">y</a> <a id="9423" class="Symbol">=</a> <a id="9425" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="9435" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9419" class="Bound">x</a> <a id="9437" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9439" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9421" class="Bound">y</a> <a id="9448" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="9451" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5624" class="Function">¬-involutive</a> <a id="9464" class="Symbol">_</a> <a id="9466" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="9472" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9474" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9476" class="Symbol">(</a><a id="9477" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9419" class="Bound">x</a> <a id="9479" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9481" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9421" class="Bound">y</a><a id="9482" class="Symbol">)</a> <a id="9485" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9489" href="Algebra.Lattice.Structures.html#5639" class="Function">¬-cong</a> <a id="9496" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="9498" href="Algebra.Lattice.Properties.BooleanAlgebra.html#8108" class="Function">¬-distribˡ-⊕</a> <a id="9511" class="Symbol">_</a> <a id="9513" class="Symbol">_</a> <a id="9515" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9521" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9523" class="Symbol">(</a><a id="9524" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9526" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9419" class="Bound">x</a> <a id="9528" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9530" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9421" class="Bound">y</a><a id="9531" class="Symbol">)</a> <a id="9534" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9538" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9159" class="Function">¬-distribʳ-⊕</a> <a id="9551" class="Symbol">_</a> <a id="9553" class="Symbol">_</a> <a id="9555" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9561" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9563" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9419" class="Bound">x</a> <a id="9565" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9567" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9569" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9421" class="Bound">y</a> <a id="9574" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.⊕-identityˡ"></a><a id="9579" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9579" class="Function">⊕-identityˡ</a> <a id="9591" class="Symbol">:</a> <a id="9593" href="Algebra.Definitions.html#1708" class="Function">LeftIdentity</a> <a id="9606" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="9608" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="9614" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9579" class="Function">⊕-identityˡ</a> <a id="9626" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9626" class="Bound">x</a> <a id="9628" class="Symbol">=</a> <a id="9630" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="9640" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="9642" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="9644" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9626" class="Bound">x</a> <a id="9661" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9664" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="9670" class="Symbol">_</a> <a id="9672" class="Symbol">_</a> <a id="9674" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9680" class="Symbol">(</a><a id="9681" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="9683" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="9685" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9626" class="Bound">x</a><a id="9686" class="Symbol">)</a> <a id="9688" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="9690" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9692" class="Symbol">(</a><a id="9693" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="9695" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="9697" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9626" class="Bound">x</a><a id="9698" class="Symbol">)</a> <a id="9701" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9704" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7488" class="Function">helper</a> <a id="9711" class="Symbol">(</a><a id="9712" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2230" class="Function">∨-identityˡ</a> <a id="9724" class="Symbol">_)</a> <a id="9727" class="Symbol">(</a><a id="9728" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2611" class="Function">∧-zeroˡ</a> <a id="9736" class="Symbol">_)</a> <a id="9739" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9745" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9626" class="Bound">x</a> <a id="9747" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="9749" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="9751" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="9766" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9769" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="9777" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5506" class="Function">⊥≉⊤</a> <a id="9781" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9787" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9626" class="Bound">x</a> <a id="9789" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="9791" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="9808" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="9811" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1746" class="Function">∧-identityʳ</a> <a id="9823" class="Symbol">_</a> <a id="9825" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="9831" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9626" class="Bound">x</a> <a id="9852" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.⊕-identityʳ"></a><a id="9857" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9857" class="Function">⊕-identityʳ</a> <a id="9869" class="Symbol">:</a> <a id="9871" href="Algebra.Definitions.html#1781" class="Function">RightIdentity</a> <a id="9885" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="9887" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="9893" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9857" class="Function">⊕-identityʳ</a> <a id="9905" class="Symbol">_</a> <a id="9907" class="Symbol">=</a> <a id="9909" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7887" class="Function">⊕-comm</a> <a id="9916" class="Symbol">_</a> <a id="9918" class="Symbol">_</a> <a id="9920" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="9922" href="Relation.Binary.Structures.html#1648" class="Function">trans</a> <a id="9928" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="9930" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9579" class="Function">⊕-identityˡ</a> <a id="9942" class="Symbol">_</a>
|
||
|
||
<a id="XorRing.⊕-identity"></a><a id="9947" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9947" class="Function">⊕-identity</a> <a id="9958" class="Symbol">:</a> <a id="9960" href="Algebra.Definitions.html#1856" class="Function">Identity</a> <a id="9969" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="9971" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="9977" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9947" class="Function">⊕-identity</a> <a id="9988" class="Symbol">=</a> <a id="9990" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9579" class="Function">⊕-identityˡ</a> <a id="10002" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="10004" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9857" class="Function">⊕-identityʳ</a>
|
||
|
||
<a id="XorRing.⊕-inverseˡ"></a><a id="10019" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10019" class="Function">⊕-inverseˡ</a> <a id="10030" class="Symbol">:</a> <a id="10032" href="Algebra.Definitions.html#2144" class="Function">LeftInverse</a> <a id="10044" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="10046" href="Function.Base.html#704" class="Function">id</a> <a id="10049" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="10055" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10019" class="Function">⊕-inverseˡ</a> <a id="10066" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10066" class="Bound">x</a> <a id="10068" class="Symbol">=</a> <a id="10070" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="10080" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10066" class="Bound">x</a> <a id="10082" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="10084" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10066" class="Bound">x</a> <a id="10100" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10103" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="10109" class="Symbol">_</a> <a id="10111" class="Symbol">_</a> <a id="10113" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="10119" class="Symbol">(</a><a id="10120" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10066" class="Bound">x</a> <a id="10122" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10124" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10066" class="Bound">x</a><a id="10125" class="Symbol">)</a> <a id="10127" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10129" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10131" class="Symbol">(</a><a id="10132" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10066" class="Bound">x</a> <a id="10134" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10136" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10066" class="Bound">x</a><a id="10137" class="Symbol">)</a> <a id="10139" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10142" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7488" class="Function">helper</a> <a id="10149" class="Symbol">(</a><a id="10150" href="Algebra.Lattice.Properties.Lattice.html#1877" class="Function">∨-idem</a> <a id="10157" class="Symbol">_)</a> <a id="10160" class="Symbol">(</a><a id="10161" href="Algebra.Lattice.Properties.Lattice.html#893" class="Function">∧-idem</a> <a id="10168" class="Symbol">_)</a> <a id="10171" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="10177" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10066" class="Bound">x</a> <a id="10179" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10181" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10183" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10066" class="Bound">x</a> <a id="10197" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10200" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="10214" class="Symbol">_</a> <a id="10216" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="10222" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="10242" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.⊕-inverseʳ"></a><a id="10247" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10247" class="Function">⊕-inverseʳ</a> <a id="10258" class="Symbol">:</a> <a id="10260" href="Algebra.Definitions.html#2232" class="Function">RightInverse</a> <a id="10273" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="10275" href="Function.Base.html#704" class="Function">id</a> <a id="10278" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="10284" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10247" class="Function">⊕-inverseʳ</a> <a id="10295" class="Symbol">_</a> <a id="10297" class="Symbol">=</a> <a id="10299" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7887" class="Function">⊕-comm</a> <a id="10306" class="Symbol">_</a> <a id="10308" class="Symbol">_</a> <a id="10310" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="10312" href="Relation.Binary.Structures.html#1648" class="Function">trans</a> <a id="10318" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="10320" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10019" class="Function">⊕-inverseˡ</a> <a id="10331" class="Symbol">_</a>
|
||
|
||
<a id="XorRing.⊕-inverse"></a><a id="10336" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10336" class="Function">⊕-inverse</a> <a id="10346" class="Symbol">:</a> <a id="10348" href="Algebra.Definitions.html#2322" class="Function">Inverse</a> <a id="10356" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="10358" href="Function.Base.html#704" class="Function">id</a> <a id="10361" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="10367" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10336" class="Function">⊕-inverse</a> <a id="10377" class="Symbol">=</a> <a id="10379" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10019" class="Function">⊕-inverseˡ</a> <a id="10390" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="10392" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10247" class="Function">⊕-inverseʳ</a>
|
||
|
||
<a id="XorRing.∧-distribˡ-⊕"></a><a id="10406" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10406" class="Function">∧-distribˡ-⊕</a> <a id="10419" class="Symbol">:</a> <a id="10421" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="10425" href="Algebra.Definitions.html#3155" class="Function Operator">DistributesOverˡ</a> <a id="10442" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="10448" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10406" class="Function">∧-distribˡ-⊕</a> <a id="10461" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10463" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10465" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a> <a id="10467" class="Symbol">=</a> <a id="10469" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="10479" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10481" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10483" class="Symbol">(</a><a id="10484" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10486" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="10488" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10489" class="Symbol">)</a> <a id="10506" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10509" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="10517" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="10519" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="10525" class="Symbol">_</a> <a id="10527" class="Symbol">_</a> <a id="10529" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="10535" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10537" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10539" class="Symbol">((</a><a id="10541" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10543" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10545" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10546" class="Symbol">)</a> <a id="10548" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10550" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10552" class="Symbol">(</a><a id="10553" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10555" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10557" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10558" class="Symbol">))</a> <a id="10562" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="10565" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="10573" class="Symbol">_</a> <a id="10575" class="Symbol">_</a> <a id="10577" class="Symbol">_</a> <a id="10579" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="10585" class="Symbol">(</a><a id="10586" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10588" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10590" class="Symbol">(</a><a id="10591" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10593" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10595" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10596" class="Symbol">))</a> <a id="10599" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10601" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10603" class="Symbol">(</a><a id="10604" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10606" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10608" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10609" class="Symbol">)</a> <a id="10612" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10615" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="10623" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="10625" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="10635" class="Symbol">_</a> <a id="10637" class="Symbol">_</a> <a id="10639" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="10645" class="Symbol">(</a><a id="10646" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10648" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10650" class="Symbol">(</a><a id="10651" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10653" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10655" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10656" class="Symbol">))</a> <a id="10659" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="10665" class="Symbol">(</a><a id="10666" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10668" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10670" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10672" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10674" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10675" class="Symbol">)</a> <a id="10692" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="10695" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2230" class="Function">∨-identityˡ</a> <a id="10707" class="Symbol">_</a> <a id="10709" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="10715" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="10717" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a>
|
||
<a id="10723" class="Symbol">((</a><a id="10725" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10727" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10729" class="Symbol">(</a><a id="10730" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10732" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10734" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10735" class="Symbol">))</a> <a id="10738" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="10744" class="Symbol">(</a><a id="10745" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10747" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10749" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10751" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10753" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10754" class="Symbol">))</a> <a id="10771" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="10774" href="Algebra.Lattice.Structures.html#4700" class="Function">∨-congʳ</a> <a id="10782" href="Algebra.Lattice.Properties.BooleanAlgebra.html#11751" class="Function">lem₃</a> <a id="10787" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="10793" class="Symbol">((</a><a id="10795" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10797" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10799" class="Symbol">(</a><a id="10800" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10802" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10804" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10805" class="Symbol">))</a> <a id="10808" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10810" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10812" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a><a id="10813" class="Symbol">)</a> <a id="10815" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a>
|
||
<a id="10821" class="Symbol">((</a><a id="10823" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10825" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10827" class="Symbol">(</a><a id="10828" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10830" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10832" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10833" class="Symbol">))</a> <a id="10836" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="10842" class="Symbol">(</a><a id="10843" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10845" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10847" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10849" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10851" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10852" class="Symbol">))</a> <a id="10869" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="10872" href="Algebra.Lattice.Structures.html#5126" class="Function">∧-distribˡ-∨</a> <a id="10885" class="Symbol">_</a> <a id="10887" class="Symbol">_</a> <a id="10889" class="Symbol">_</a> <a id="10891" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="10897" class="Symbol">(</a><a id="10898" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10900" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10902" class="Symbol">(</a><a id="10903" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10905" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10907" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10908" class="Symbol">))</a> <a id="10911" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="10917" class="Symbol">(</a><a id="10918" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10920" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10922" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10924" class="Symbol">(</a><a id="10925" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10927" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10929" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10931" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="10933" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10934" class="Symbol">))</a> <a id="10944" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="10947" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="10955" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="10957" href="Algebra.Lattice.Structures.html#4639" class="Function">∨-congˡ</a> <a id="10965" class="Symbol">(</a><a id="10966" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="10976" class="Symbol">_</a> <a id="10978" class="Symbol">_)</a> <a id="10981" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="10987" class="Symbol">(</a><a id="10988" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="10990" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="10992" class="Symbol">(</a><a id="10993" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="10995" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="10997" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="10998" class="Symbol">))</a> <a id="11001" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="11007" class="Symbol">(</a><a id="11008" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="11010" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11012" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="11014" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="11016" class="Symbol">(</a><a id="11017" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11019" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11021" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11022" class="Symbol">))</a> <a id="11034" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="11037" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="11045" class="Symbol">(</a><a id="11046" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="11056" class="Symbol">_</a> <a id="11058" class="Symbol">_)</a> <a id="11061" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="11067" class="Symbol">(</a><a id="11068" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11070" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11072" class="Symbol">(</a><a id="11073" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11075" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="11077" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11078" class="Symbol">))</a> <a id="11081" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="11087" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="11089" class="Symbol">(</a><a id="11090" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11092" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11094" class="Symbol">(</a><a id="11095" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11097" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11099" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11100" class="Symbol">))</a> <a id="11114" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11117" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7488" class="Function">helper</a> <a id="11124" href="Relation.Binary.Structures.html#1596" class="Function">refl</a> <a id="11129" href="Algebra.Lattice.Properties.BooleanAlgebra.html#11514" class="Function">lem₁</a> <a id="11134" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="11140" class="Symbol">(</a><a id="11141" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11143" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11145" class="Symbol">(</a><a id="11146" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11148" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="11150" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11151" class="Symbol">))</a> <a id="11154" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="11160" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="11162" class="Symbol">((</a><a id="11164" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11166" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11168" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a><a id="11169" class="Symbol">)</a> <a id="11171" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11173" class="Symbol">(</a><a id="11174" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11176" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11178" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11179" class="Symbol">))</a> <a id="11187" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11190" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="11198" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="11200" href="Algebra.Lattice.Structures.html#5126" class="Function">∧-distribˡ-∨</a> <a id="11213" class="Symbol">_</a> <a id="11215" class="Symbol">_</a> <a id="11217" class="Symbol">_</a> <a id="11219" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="11225" class="Symbol">((</a><a id="11227" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11229" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11231" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a><a id="11232" class="Symbol">)</a> <a id="11234" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="11236" class="Symbol">(</a><a id="11237" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11239" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11241" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11242" class="Symbol">))</a> <a id="11245" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="11251" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="11253" class="Symbol">((</a><a id="11255" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11257" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11259" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a><a id="11260" class="Symbol">)</a> <a id="11262" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11264" class="Symbol">(</a><a id="11265" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11267" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11269" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11270" class="Symbol">))</a> <a id="11278" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="11281" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="11287" class="Symbol">_</a> <a id="11289" class="Symbol">_</a> <a id="11291" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="11297" class="Symbol">(</a><a id="11298" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11300" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11302" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a><a id="11303" class="Symbol">)</a> <a id="11305" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="11307" class="Symbol">(</a><a id="11308" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11310" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11312" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11313" class="Symbol">)</a> <a id="11324" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
<a id="11332" class="Keyword">where</a>
|
||
<a id="11344" href="Algebra.Lattice.Properties.BooleanAlgebra.html#11344" class="Function">lem₂</a> <a id="11349" class="Symbol">=</a> <a id="11351" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="11365" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11367" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11369" class="Symbol">(</a><a id="11370" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11372" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11374" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11375" class="Symbol">)</a> <a id="11378" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="11381" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="11389" class="Symbol">_</a> <a id="11391" class="Symbol">_</a> <a id="11393" class="Symbol">_</a> <a id="11395" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="11405" class="Symbol">(</a><a id="11406" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11408" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11410" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a><a id="11411" class="Symbol">)</a> <a id="11413" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11415" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a> <a id="11418" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11421" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="11429" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="11431" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="11438" class="Symbol">_</a> <a id="11440" class="Symbol">_</a> <a id="11442" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="11452" class="Symbol">(</a><a id="11453" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11455" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11457" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a><a id="11458" class="Symbol">)</a> <a id="11460" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11462" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a> <a id="11465" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11468" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="11476" class="Symbol">_</a> <a id="11478" class="Symbol">_</a> <a id="11480" class="Symbol">_</a> <a id="11482" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="11492" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11494" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11496" class="Symbol">(</a><a id="11497" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11499" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11501" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11502" class="Symbol">)</a> <a id="11505" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="11514" href="Algebra.Lattice.Properties.BooleanAlgebra.html#11514" class="Function">lem₁</a> <a id="11519" class="Symbol">=</a> <a id="11521" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="11535" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11537" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11539" class="Symbol">(</a><a id="11540" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11542" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11544" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11545" class="Symbol">)</a> <a id="11554" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="11557" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="11565" class="Symbol">(</a><a id="11566" href="Algebra.Lattice.Properties.Lattice.html#893" class="Function">∧-idem</a> <a id="11573" class="Symbol">_)</a> <a id="11576" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="11586" class="Symbol">(</a><a id="11587" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11589" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11591" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a><a id="11592" class="Symbol">)</a> <a id="11594" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11596" class="Symbol">(</a><a id="11597" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11599" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11601" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11602" class="Symbol">)</a> <a id="11605" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11608" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="11616" class="Symbol">_</a> <a id="11618" class="Symbol">_</a> <a id="11620" class="Symbol">_</a> <a id="11622" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="11632" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11634" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11636" class="Symbol">(</a><a id="11637" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11639" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11641" class="Symbol">(</a><a id="11642" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11644" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11646" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11647" class="Symbol">))</a> <a id="11651" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11654" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="11662" href="Algebra.Lattice.Properties.BooleanAlgebra.html#11344" class="Function">lem₂</a> <a id="11667" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="11677" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11679" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11681" class="Symbol">(</a><a id="11682" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11684" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11686" class="Symbol">(</a><a id="11687" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11689" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11691" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11692" class="Symbol">))</a> <a id="11696" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="11699" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="11707" class="Symbol">_</a> <a id="11709" class="Symbol">_</a> <a id="11711" class="Symbol">_</a> <a id="11713" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="11723" class="Symbol">(</a><a id="11724" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11726" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11728" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a><a id="11729" class="Symbol">)</a> <a id="11731" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11733" class="Symbol">(</a><a id="11734" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11736" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11738" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11739" class="Symbol">)</a> <a id="11742" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="11751" href="Algebra.Lattice.Properties.BooleanAlgebra.html#11751" class="Function">lem₃</a> <a id="11756" class="Symbol">=</a> <a id="11758" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="11772" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="11795" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="11798" href="Algebra.Lattice.Properties.BooleanAlgebra.html#2378" class="Function">∧-zeroʳ</a> <a id="11806" class="Symbol">_</a> <a id="11808" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="11818" class="Symbol">(</a><a id="11819" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11821" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="11823" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11824" class="Symbol">)</a> <a id="11826" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11828" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="11841" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="11844" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="11852" class="Symbol">(</a><a id="11853" href="Algebra.Lattice.Structures.html#5961" class="Function">∧-complementʳ</a> <a id="11867" class="Symbol">_)</a> <a id="11870" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="11880" class="Symbol">(</a><a id="11881" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11883" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="11885" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11886" class="Symbol">)</a> <a id="11888" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11890" class="Symbol">(</a><a id="11891" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11893" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11895" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="11897" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a><a id="11898" class="Symbol">)</a> <a id="11903" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="11906" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="11914" class="Symbol">_</a> <a id="11916" class="Symbol">_</a> <a id="11918" class="Symbol">_</a> <a id="11920" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="11930" class="Symbol">((</a><a id="11932" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11934" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="11936" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11937" class="Symbol">)</a> <a id="11939" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11941" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a><a id="11942" class="Symbol">)</a> <a id="11944" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11946" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="11948" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11953" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="11956" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="11964" class="Symbol">(</a><a id="11965" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="11972" class="Symbol">_</a> <a id="11974" class="Symbol">_)</a> <a id="11977" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="11987" class="Symbol">(</a><a id="11988" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="11990" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="11992" class="Symbol">(</a><a id="11993" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10463" class="Bound">y</a> <a id="11995" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="11997" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10465" class="Bound">z</a><a id="11998" class="Symbol">))</a> <a id="12001" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12003" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="12005" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10461" class="Bound">x</a> <a id="12010" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.∧-distribʳ-⊕"></a><a id="12015" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12015" class="Function">∧-distribʳ-⊕</a> <a id="12028" class="Symbol">:</a> <a id="12030" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="12034" href="Algebra.Definitions.html#3274" class="Function Operator">DistributesOverʳ</a> <a id="12051" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="12057" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12015" class="Function">∧-distribʳ-⊕</a> <a id="12070" class="Symbol">=</a> <a id="12072" href="Algebra.Consequences.Setoid.html#7938" class="Function">comm∧distrˡ⇒distrʳ</a> <a id="12091" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7595" class="Function">⊕-cong</a> <a id="12098" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="12105" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10406" class="Function">∧-distribˡ-⊕</a>
|
||
|
||
<a id="XorRing.∧-distrib-⊕"></a><a id="12121" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12121" class="Function">∧-distrib-⊕</a> <a id="12133" class="Symbol">:</a> <a id="12135" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="12139" href="Algebra.Definitions.html#3393" class="Function Operator">DistributesOver</a> <a id="12155" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="12161" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12121" class="Function">∧-distrib-⊕</a> <a id="12173" class="Symbol">=</a> <a id="12175" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10406" class="Function">∧-distribˡ-⊕</a> <a id="12188" href="Agda.Builtin.Sigma.html#235" class="InductiveConstructor Operator">,</a> <a id="12190" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12015" class="Function">∧-distribʳ-⊕</a>
|
||
|
||
<a id="12206" class="Keyword">private</a>
|
||
|
||
<a id="XorRing.lemma₂"></a><a id="12219" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12219" class="Function">lemma₂</a> <a id="12226" class="Symbol">:</a> <a id="12228" class="Symbol">∀</a> <a id="12230" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12230" class="Bound">x</a> <a id="12232" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12232" class="Bound">y</a> <a id="12234" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12234" class="Bound">u</a> <a id="12236" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12236" class="Bound">v</a> <a id="12238" class="Symbol">→</a>
|
||
<a id="12253" class="Symbol">(</a><a id="12254" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12230" class="Bound">x</a> <a id="12256" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12258" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12232" class="Bound">y</a><a id="12259" class="Symbol">)</a> <a id="12261" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12263" class="Symbol">(</a><a id="12264" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12234" class="Bound">u</a> <a id="12266" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12268" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12236" class="Bound">v</a><a id="12269" class="Symbol">)</a> <a id="12271" href="Algebra.Lattice.Bundles.html#5699" class="Field Operator">≈</a>
|
||
<a id="12286" class="Symbol">((</a><a id="12288" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12230" class="Bound">x</a> <a id="12290" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12292" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12234" class="Bound">u</a><a id="12293" class="Symbol">)</a> <a id="12295" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12297" class="Symbol">(</a><a id="12298" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12232" class="Bound">y</a> <a id="12300" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12302" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12234" class="Bound">u</a><a id="12303" class="Symbol">))</a> <a id="12306" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="12321" class="Symbol">((</a><a id="12323" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12230" class="Bound">x</a> <a id="12325" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12327" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12236" class="Bound">v</a><a id="12328" class="Symbol">)</a> <a id="12330" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12332" class="Symbol">(</a><a id="12333" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12232" class="Bound">y</a> <a id="12335" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12337" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12236" class="Bound">v</a><a id="12338" class="Symbol">))</a>
|
||
<a id="12345" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12219" class="Function">lemma₂</a> <a id="12352" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12352" class="Bound">x</a> <a id="12354" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12354" class="Bound">y</a> <a id="12356" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12356" class="Bound">u</a> <a id="12358" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12358" class="Bound">v</a> <a id="12360" class="Symbol">=</a> <a id="12362" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="12376" class="Symbol">(</a><a id="12377" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12352" class="Bound">x</a> <a id="12379" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12381" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12354" class="Bound">y</a><a id="12382" class="Symbol">)</a> <a id="12384" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12386" class="Symbol">(</a><a id="12387" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12356" class="Bound">u</a> <a id="12389" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12391" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12358" class="Bound">v</a><a id="12392" class="Symbol">)</a> <a id="12407" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="12410" href="Algebra.Lattice.Structures.html#4978" class="Function">∨-distribˡ-∧</a> <a id="12423" class="Symbol">_</a> <a id="12425" class="Symbol">_</a> <a id="12427" class="Symbol">_</a> <a id="12429" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="12439" class="Symbol">((</a><a id="12441" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12352" class="Bound">x</a> <a id="12443" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12445" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12354" class="Bound">y</a><a id="12446" class="Symbol">)</a> <a id="12448" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12450" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12356" class="Bound">u</a><a id="12451" class="Symbol">)</a> <a id="12453" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12455" class="Symbol">((</a><a id="12457" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12352" class="Bound">x</a> <a id="12459" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12461" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12354" class="Bound">y</a><a id="12462" class="Symbol">)</a> <a id="12464" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12466" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12358" class="Bound">v</a><a id="12467" class="Symbol">)</a> <a id="12470" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="12473" href="Algebra.Lattice.Structures.html#5052" class="Function">∨-distribʳ-∧</a> <a id="12486" class="Symbol">_</a> <a id="12488" class="Symbol">_</a> <a id="12490" class="Symbol">_</a>
|
||
<a id="12536" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="12538" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="12545" href="Function.Base.html#4322" class="Function Operator">⟩</a>
|
||
<a id="12589" href="Algebra.Lattice.Structures.html#5052" class="Function">∨-distribʳ-∧</a> <a id="12602" class="Symbol">_</a> <a id="12604" class="Symbol">_</a> <a id="12606" class="Symbol">_</a> <a id="12608" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="12618" class="Symbol">((</a><a id="12620" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12352" class="Bound">x</a> <a id="12622" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12624" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12356" class="Bound">u</a><a id="12625" class="Symbol">)</a> <a id="12627" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12629" class="Symbol">(</a><a id="12630" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12354" class="Bound">y</a> <a id="12632" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12634" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12356" class="Bound">u</a><a id="12635" class="Symbol">))</a> <a id="12638" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="12648" class="Symbol">((</a><a id="12650" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12352" class="Bound">x</a> <a id="12652" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12654" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12358" class="Bound">v</a><a id="12655" class="Symbol">)</a> <a id="12657" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12659" class="Symbol">(</a><a id="12660" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12354" class="Bound">y</a> <a id="12662" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12664" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12358" class="Bound">v</a><a id="12665" class="Symbol">))</a> <a id="12679" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.⊕-assoc"></a><a id="12684" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12684" class="Function">⊕-assoc</a> <a id="12692" class="Symbol">:</a> <a id="12694" href="Algebra.Definitions.html#1548" class="Function">Associative</a> <a id="12706" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="12712" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12684" class="Function">⊕-assoc</a> <a id="12720" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="12722" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="12724" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a> <a id="12726" class="Symbol">=</a> <a id="12728" href="Relation.Binary.Structures.html#1622" class="Function">sym</a> <a id="12732" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="12734" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="12744" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="12746" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="12748" class="Symbol">(</a><a id="12749" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="12751" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="12753" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="12754" class="Symbol">)</a> <a id="12787" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="12790" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7595" class="Function">⊕-cong</a> <a id="12797" href="Relation.Binary.Structures.html#1596" class="Function">refl</a> <a id="12802" class="Symbol">(</a><a id="12803" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="12809" class="Symbol">_</a> <a id="12811" class="Symbol">_)</a> <a id="12814" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="12820" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="12822" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="12824" class="Symbol">((</a><a id="12826" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="12828" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12830" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="12831" class="Symbol">)</a> <a id="12833" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12835" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="12837" class="Symbol">(</a><a id="12838" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="12840" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12842" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="12843" class="Symbol">))</a> <a id="12863" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="12866" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="12872" class="Symbol">_</a> <a id="12874" class="Symbol">_</a> <a id="12876" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="12884" class="Symbol">(</a><a id="12885" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="12887" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12889" class="Symbol">((</a><a id="12891" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="12893" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12895" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="12896" class="Symbol">)</a> <a id="12898" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12900" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="12902" class="Symbol">(</a><a id="12903" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="12905" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12907" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="12908" class="Symbol">)))</a> <a id="12912" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="12918" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="12920" class="Symbol">(</a><a id="12921" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="12923" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12925" class="Symbol">((</a><a id="12927" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="12929" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12931" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="12932" class="Symbol">)</a> <a id="12934" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12936" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="12938" class="Symbol">(</a><a id="12939" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="12941" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="12943" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="12944" class="Symbol">)))</a> <a id="12961" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="12964" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="12971" href="Algebra.Lattice.Properties.BooleanAlgebra.html#14727" class="Function">lem₃</a> <a id="12976" href="Algebra.Lattice.Properties.BooleanAlgebra.html#15578" class="Function">lem₄</a> <a id="12981" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="12987" class="Symbol">(((</a><a id="12990" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="12992" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12994" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="12995" class="Symbol">)</a> <a id="12997" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="12999" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13000" class="Symbol">)</a> <a id="13002" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13004" class="Symbol">((</a><a id="13006" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13008" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13010" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13012" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13013" class="Symbol">)</a> <a id="13015" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13017" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13019" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13020" class="Symbol">))</a> <a id="13023" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="13029" class="Symbol">(((</a><a id="13032" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13034" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13036" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13038" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13040" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13041" class="Symbol">)</a> <a id="13043" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13045" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13046" class="Symbol">)</a> <a id="13048" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13050" class="Symbol">((</a><a id="13052" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13054" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13056" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13058" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13059" class="Symbol">)</a> <a id="13061" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13063" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13065" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13066" class="Symbol">))</a> <a id="13072" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="13075" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="13083" class="Symbol">_</a> <a id="13085" class="Symbol">_</a> <a id="13087" class="Symbol">_</a> <a id="13089" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="13095" class="Symbol">((</a><a id="13097" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13099" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13101" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13102" class="Symbol">)</a> <a id="13104" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13106" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13107" class="Symbol">)</a> <a id="13109" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="13115" class="Symbol">(((</a><a id="13118" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13120" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13122" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13124" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13125" class="Symbol">)</a> <a id="13127" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13129" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13131" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13132" class="Symbol">)</a> <a id="13134" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="13141" class="Symbol">(((</a><a id="13144" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13146" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13148" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13150" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13152" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13153" class="Symbol">)</a> <a id="13155" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13157" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13158" class="Symbol">)</a> <a id="13160" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13162" class="Symbol">((</a><a id="13164" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13166" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13168" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13170" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13171" class="Symbol">)</a> <a id="13173" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13175" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13177" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13178" class="Symbol">)))</a> <a id="13183" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="13186" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="13194" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16028" class="Function">lem₅</a> <a id="13199" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="13205" class="Symbol">((</a><a id="13207" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13209" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13211" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13212" class="Symbol">)</a> <a id="13214" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13216" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13217" class="Symbol">)</a> <a id="13219" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="13225" class="Symbol">(((</a><a id="13228" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13230" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13232" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13234" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13236" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13237" class="Symbol">)</a> <a id="13239" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13241" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13242" class="Symbol">)</a> <a id="13244" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="13251" class="Symbol">(((</a><a id="13254" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13256" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13258" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13260" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13261" class="Symbol">)</a> <a id="13263" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13265" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13267" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13268" class="Symbol">)</a> <a id="13270" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13272" class="Symbol">((</a><a id="13274" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13276" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13278" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13280" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13281" class="Symbol">)</a> <a id="13283" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13285" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13287" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13288" class="Symbol">)))</a> <a id="13293" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="13296" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="13304" class="Symbol">_</a> <a id="13306" class="Symbol">_</a> <a id="13308" class="Symbol">_</a> <a id="13310" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="13316" class="Symbol">(((</a><a id="13319" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13321" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13323" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13324" class="Symbol">)</a> <a id="13326" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13328" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13329" class="Symbol">)</a> <a id="13331" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13333" class="Symbol">((</a><a id="13335" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13337" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13339" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13341" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13343" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13344" class="Symbol">)</a> <a id="13346" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13348" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13349" class="Symbol">))</a> <a id="13352" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="13358" class="Symbol">(((</a><a id="13361" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13363" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13365" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13367" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13368" class="Symbol">)</a> <a id="13370" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13372" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13374" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13375" class="Symbol">)</a> <a id="13377" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13379" class="Symbol">((</a><a id="13381" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13383" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13385" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13387" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13388" class="Symbol">)</a> <a id="13390" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13392" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13394" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13395" class="Symbol">))</a> <a id="13401" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="13404" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="13411" href="Algebra.Lattice.Properties.BooleanAlgebra.html#13660" class="Function">lem₁</a> <a id="13416" href="Algebra.Lattice.Properties.BooleanAlgebra.html#14465" class="Function">lem₂</a> <a id="13421" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="13429" class="Symbol">(((</a><a id="13432" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13434" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13436" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13437" class="Symbol">)</a> <a id="13439" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13441" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13443" class="Symbol">(</a><a id="13444" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13446" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13448" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13449" class="Symbol">))</a> <a id="13452" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13454" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13455" class="Symbol">)</a> <a id="13457" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="13463" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13465" class="Symbol">(((</a><a id="13468" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13470" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13472" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13473" class="Symbol">)</a> <a id="13475" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13477" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13479" class="Symbol">(</a><a id="13480" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13482" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13484" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13485" class="Symbol">))</a> <a id="13488" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13490" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13491" class="Symbol">)</a> <a id="13506" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="13509" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="13515" class="Symbol">_</a> <a id="13517" class="Symbol">_</a> <a id="13519" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="13525" class="Symbol">((</a><a id="13527" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13529" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13531" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13532" class="Symbol">)</a> <a id="13534" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13536" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13538" class="Symbol">(</a><a id="13539" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13541" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13543" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13544" class="Symbol">))</a> <a id="13547" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="13549" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a> <a id="13568" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="13571" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7595" class="Function">⊕-cong</a> <a id="13578" class="Symbol">(</a><a id="13579" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7363" class="Bound">⊕-def</a> <a id="13585" class="Symbol">_</a> <a id="13587" class="Symbol">_)</a> <a id="13590" href="Relation.Binary.Structures.html#1596" class="Function">refl</a> <a id="13595" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="13601" class="Symbol">(</a><a id="13602" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13604" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="13606" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13607" class="Symbol">)</a> <a id="13609" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">⊕</a> <a id="13611" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a> <a id="13644" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
<a id="13650" class="Keyword">where</a>
|
||
<a id="13660" href="Algebra.Lattice.Properties.BooleanAlgebra.html#13660" class="Function">lem₁</a> <a id="13665" class="Symbol">=</a> <a id="13667" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="13679" class="Symbol">((</a><a id="13681" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13683" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13685" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13686" class="Symbol">)</a> <a id="13688" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13690" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13691" class="Symbol">)</a> <a id="13693" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13695" class="Symbol">((</a><a id="13697" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13699" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13701" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13703" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13705" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13706" class="Symbol">)</a> <a id="13708" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13710" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="13711" class="Symbol">)</a> <a id="13714" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="13717" href="Algebra.Lattice.Structures.html#5052" class="Function">∨-distribʳ-∧</a> <a id="13730" class="Symbol">_</a> <a id="13732" class="Symbol">_</a> <a id="13734" class="Symbol">_</a> <a id="13736" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="13744" class="Symbol">((</a><a id="13746" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13748" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13750" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13751" class="Symbol">)</a> <a id="13753" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13755" class="Symbol">(</a><a id="13756" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13758" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13760" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13762" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13764" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13765" class="Symbol">))</a> <a id="13768" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13770" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a> <a id="13779" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="13782" href="Algebra.Lattice.Structures.html#4700" class="Function">∨-congʳ</a> <a id="13790" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="13792" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="13800" class="Symbol">(</a><a id="13801" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="13811" class="Symbol">_</a> <a id="13813" class="Symbol">_)</a> <a id="13816" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="13824" class="Symbol">((</a><a id="13826" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13828" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13830" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13831" class="Symbol">)</a> <a id="13833" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13835" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13837" class="Symbol">(</a><a id="13838" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13840" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13842" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13843" class="Symbol">))</a> <a id="13846" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13848" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a> <a id="13859" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="13866" href="Algebra.Lattice.Properties.BooleanAlgebra.html#13866" class="Function">lem₂′</a> <a id="13872" class="Symbol">=</a> <a id="13874" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="13886" class="Symbol">(</a><a id="13887" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13889" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13891" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13893" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13894" class="Symbol">)</a> <a id="13896" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13898" class="Symbol">(</a><a id="13899" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13901" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13903" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13905" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13906" class="Symbol">)</a> <a id="13921" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="13924" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="13931" class="Symbol">(</a><a id="13932" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1914" class="Function">∧-identityˡ</a> <a id="13944" class="Symbol">_)</a> <a id="13947" class="Symbol">(</a><a id="13948" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1746" class="Function">∧-identityʳ</a> <a id="13960" class="Symbol">_)</a> <a id="13963" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="13971" class="Symbol">(</a><a id="13972" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="13974" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13976" class="Symbol">(</a><a id="13977" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13979" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13981" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13983" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13984" class="Symbol">))</a> <a id="13987" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="13989" class="Symbol">((</a><a id="13991" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="13993" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="13995" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="13997" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="13998" class="Symbol">)</a> <a id="14000" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14002" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a><a id="14003" class="Symbol">)</a> <a id="14006" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="14009" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a>
|
||
<a id="14062" class="Symbol">(</a><a id="14063" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="14070" class="Symbol">(</a><a id="14071" href="Algebra.Lattice.Structures.html#5738" class="Function">∨-complementˡ</a> <a id="14085" class="Symbol">_)</a> <a id="14088" class="Symbol">(</a><a id="14089" href="Algebra.Lattice.Structures.html#4113" class="Function">∨-comm</a> <a id="14096" class="Symbol">_</a> <a id="14098" class="Symbol">_))</a>
|
||
<a id="14148" class="Symbol">(</a><a id="14149" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="14157" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="14159" href="Algebra.Lattice.Structures.html#5738" class="Function">∨-complementˡ</a> <a id="14173" class="Symbol">_)</a> <a id="14176" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="14184" class="Symbol">((</a><a id="14186" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14188" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14190" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14192" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a><a id="14193" class="Symbol">)</a> <a id="14195" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14197" class="Symbol">(</a><a id="14198" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14200" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="14202" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14204" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a><a id="14205" class="Symbol">))</a> <a id="14208" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="14216" class="Symbol">((</a><a id="14218" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14220" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14222" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14224" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14225" class="Symbol">)</a> <a id="14227" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14229" class="Symbol">(</a><a id="14230" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14232" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="14234" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14236" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14237" class="Symbol">))</a> <a id="14251" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="14254" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12219" class="Function">lemma₂</a> <a id="14261" class="Symbol">_</a> <a id="14263" class="Symbol">_</a> <a id="14265" class="Symbol">_</a> <a id="14267" class="Symbol">_</a> <a id="14269" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="14277" class="Symbol">(</a><a id="14278" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14280" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14282" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14284" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14286" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14287" class="Symbol">)</a> <a id="14289" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14291" class="Symbol">(</a><a id="14292" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14294" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14296" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14297" class="Symbol">)</a> <a id="14312" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="14315" href="Algebra.Lattice.Structures.html#4181" class="Function">∨-cong</a> <a id="14322" class="Symbol">(</a><a id="14323" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6880" class="Function">deMorgan₂</a> <a id="14333" class="Symbol">_</a> <a id="14335" class="Symbol">_)</a> <a id="14338" class="Symbol">(</a><a id="14339" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5624" class="Function">¬-involutive</a> <a id="14352" class="Symbol">_)</a> <a id="14355" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="14363" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14365" class="Symbol">(</a><a id="14366" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14368" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14370" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14371" class="Symbol">)</a> <a id="14373" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14375" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14377" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14379" class="Symbol">(</a><a id="14380" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14382" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14384" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14385" class="Symbol">)</a> <a id="14398" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="14401" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="14411" class="Symbol">_</a> <a id="14413" class="Symbol">_</a> <a id="14415" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="14423" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14425" class="Symbol">((</a><a id="14427" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14429" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14431" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14432" class="Symbol">)</a> <a id="14434" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14436" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14438" class="Symbol">(</a><a id="14439" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14441" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14443" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14444" class="Symbol">))</a> <a id="14458" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="14465" href="Algebra.Lattice.Properties.BooleanAlgebra.html#14465" class="Function">lem₂</a> <a id="14470" class="Symbol">=</a> <a id="14472" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="14484" class="Symbol">((</a><a id="14486" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14488" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14490" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14492" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14493" class="Symbol">)</a> <a id="14495" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14497" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14499" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14500" class="Symbol">)</a> <a id="14502" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14504" class="Symbol">((</a><a id="14506" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14508" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14510" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14512" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14513" class="Symbol">)</a> <a id="14515" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14517" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14519" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14520" class="Symbol">)</a> <a id="14523" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="14526" href="Algebra.Lattice.Structures.html#5052" class="Function">∨-distribʳ-∧</a> <a id="14539" class="Symbol">_</a> <a id="14541" class="Symbol">_</a> <a id="14543" class="Symbol">_</a> <a id="14545" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="14553" class="Symbol">((</a><a id="14555" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14557" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14559" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14561" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14562" class="Symbol">)</a> <a id="14564" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14566" class="Symbol">(</a><a id="14567" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14569" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14571" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14573" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14574" class="Symbol">))</a> <a id="14577" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14579" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14581" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a> <a id="14592" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="14595" href="Algebra.Lattice.Structures.html#4700" class="Function">∨-congʳ</a> <a id="14603" href="Algebra.Lattice.Properties.BooleanAlgebra.html#13866" class="Function">lem₂′</a> <a id="14609" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="14617" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14619" class="Symbol">((</a><a id="14621" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14623" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14625" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14626" class="Symbol">)</a> <a id="14628" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14630" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14632" class="Symbol">(</a><a id="14633" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14635" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14637" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14638" class="Symbol">))</a> <a id="14641" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14643" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14645" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a> <a id="14656" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="14659" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="14669" class="Symbol">_</a> <a id="14671" class="Symbol">_</a> <a id="14673" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="14681" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14683" class="Symbol">(((</a><a id="14686" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14688" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14690" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14691" class="Symbol">)</a> <a id="14693" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14695" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14697" class="Symbol">(</a><a id="14698" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14700" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14702" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14703" class="Symbol">))</a> <a id="14706" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14708" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14709" class="Symbol">)</a> <a id="14720" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="14727" href="Algebra.Lattice.Properties.BooleanAlgebra.html#14727" class="Function">lem₃</a> <a id="14732" class="Symbol">=</a> <a id="14734" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="14746" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14748" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14750" class="Symbol">((</a><a id="14752" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="14754" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14756" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14757" class="Symbol">)</a> <a id="14759" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14761" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14763" class="Symbol">(</a><a id="14764" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="14766" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14768" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14769" class="Symbol">))</a> <a id="14781" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="14784" href="Algebra.Lattice.Structures.html#4639" class="Function">∨-congˡ</a> <a id="14792" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="14794" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="14802" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="14804" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="14814" class="Symbol">_</a> <a id="14816" class="Symbol">_</a> <a id="14818" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="14826" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14828" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14830" class="Symbol">((</a><a id="14832" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="14834" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14836" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14837" class="Symbol">)</a> <a id="14839" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14841" class="Symbol">(</a><a id="14842" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14844" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="14846" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14848" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14850" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14851" class="Symbol">))</a> <a id="14861" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="14864" href="Algebra.Lattice.Structures.html#4978" class="Function">∨-distribˡ-∧</a> <a id="14877" class="Symbol">_</a> <a id="14879" class="Symbol">_</a> <a id="14881" class="Symbol">_</a> <a id="14883" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="14891" class="Symbol">(</a><a id="14892" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14894" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14896" class="Symbol">(</a><a id="14897" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="14899" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14901" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14902" class="Symbol">))</a> <a id="14905" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14907" class="Symbol">(</a><a id="14908" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14910" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14912" class="Symbol">(</a><a id="14913" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14915" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="14917" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14919" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="14921" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14922" class="Symbol">))</a> <a id="14926" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="14929" href="Algebra.Lattice.Structures.html#4147" class="Function">∨-assoc</a> <a id="14937" class="Symbol">_</a> <a id="14939" class="Symbol">_</a> <a id="14941" class="Symbol">_</a> <a id="14943" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="14945" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="14952" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="14954" href="Algebra.Lattice.Structures.html#4147" class="Function">∨-assoc</a> <a id="14962" class="Symbol">_</a> <a id="14964" class="Symbol">_</a> <a id="14966" class="Symbol">_</a> <a id="14968" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="14976" class="Symbol">((</a><a id="14978" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14980" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14982" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="14983" class="Symbol">)</a> <a id="14985" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14987" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="14988" class="Symbol">)</a> <a id="14990" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="14992" class="Symbol">((</a><a id="14994" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="14996" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="14998" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15000" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="15001" class="Symbol">)</a> <a id="15003" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15005" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15007" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15008" class="Symbol">)</a> <a id="15011" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="15018" href="Algebra.Lattice.Properties.BooleanAlgebra.html#15018" class="Function">lem₄′</a> <a id="15024" class="Symbol">=</a> <a id="15026" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="15038" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15040" class="Symbol">((</a><a id="15042" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15044" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15046" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15047" class="Symbol">)</a> <a id="15049" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15051" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15053" class="Symbol">(</a><a id="15054" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15056" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15058" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15059" class="Symbol">))</a> <a id="15065" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="15068" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="15078" class="Symbol">_</a> <a id="15080" class="Symbol">_</a> <a id="15082" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="15090" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15092" class="Symbol">(</a><a id="15093" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15095" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15097" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15098" class="Symbol">)</a> <a id="15100" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15102" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15104" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15106" class="Symbol">(</a><a id="15107" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15109" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15111" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15112" class="Symbol">)</a> <a id="15117" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="15120" href="Algebra.Lattice.Properties.BooleanAlgebra.html#6880" class="Function">deMorgan₂</a> <a id="15130" class="Symbol">_</a> <a id="15132" class="Symbol">_</a> <a id="15134" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="15136" href="Algebra.Lattice.Structures.html#4181" class="Function">∨-cong</a> <a id="15143" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="15145" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5624" class="Function">¬-involutive</a> <a id="15158" class="Symbol">_</a> <a id="15160" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="15168" class="Symbol">(</a><a id="15169" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15171" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15173" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15175" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15177" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15178" class="Symbol">)</a> <a id="15180" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15182" class="Symbol">(</a><a id="15183" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15185" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15187" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15188" class="Symbol">)</a> <a id="15195" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="15198" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12219" class="Function">lemma₂</a> <a id="15205" class="Symbol">_</a> <a id="15207" class="Symbol">_</a> <a id="15209" class="Symbol">_</a> <a id="15211" class="Symbol">_</a> <a id="15213" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="15221" class="Symbol">((</a><a id="15223" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15225" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15227" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15229" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="15230" class="Symbol">)</a> <a id="15232" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15234" class="Symbol">(</a><a id="15235" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15237" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a> <a id="15239" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15241" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="15242" class="Symbol">))</a> <a id="15245" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="15253" class="Symbol">((</a><a id="15255" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15257" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15259" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15261" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15262" class="Symbol">)</a> <a id="15264" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15266" class="Symbol">(</a><a id="15267" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15269" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a> <a id="15271" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15273" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15274" class="Symbol">))</a> <a id="15280" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="15283" class="Symbol">(</a><a id="15284" href="Algebra.Lattice.Structures.html#5738" class="Function">∨-complementˡ</a> <a id="15298" class="Symbol">_</a> <a id="15300" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="15302" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="15309" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="15311" href="Algebra.Lattice.Structures.html#4113" class="Function">∨-comm</a> <a id="15318" class="Symbol">_</a> <a id="15320" class="Symbol">_)</a>
|
||
<a id="15361" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="15363" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="15370" href="Function.Base.html#4322" class="Function Operator">⟩</a>
|
||
<a id="15407" class="Symbol">(</a><a id="15408" href="Algebra.Lattice.Structures.html#4516" class="Function">∧-congˡ</a> <a id="15416" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="15418" href="Algebra.Lattice.Structures.html#5738" class="Function">∨-complementˡ</a> <a id="15432" class="Symbol">_)</a> <a id="15435" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="15443" class="Symbol">(</a><a id="15444" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a> <a id="15446" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15448" class="Symbol">(</a><a id="15449" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15451" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15453" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15455" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15456" class="Symbol">))</a> <a id="15459" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="15467" class="Symbol">((</a><a id="15469" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15471" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15473" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15475" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15476" class="Symbol">)</a> <a id="15478" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15480" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a><a id="15481" class="Symbol">)</a> <a id="15494" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="15497" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1914" class="Function">∧-identityˡ</a> <a id="15509" class="Symbol">_</a> <a id="15511" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="15513" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="15520" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="15522" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1746" class="Function">∧-identityʳ</a> <a id="15534" class="Symbol">_</a> <a id="15536" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="15544" class="Symbol">(</a><a id="15545" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15547" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15549" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15551" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15552" class="Symbol">)</a> <a id="15554" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15556" class="Symbol">(</a><a id="15557" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15559" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15561" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15563" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15564" class="Symbol">)</a> <a id="15571" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="15578" href="Algebra.Lattice.Properties.BooleanAlgebra.html#15578" class="Function">lem₄</a> <a id="15583" class="Symbol">=</a> <a id="15585" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="15597" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15599" class="Symbol">(</a><a id="15600" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="15602" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15604" class="Symbol">((</a><a id="15606" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15608" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15610" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15611" class="Symbol">)</a> <a id="15613" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15615" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15617" class="Symbol">(</a><a id="15618" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15620" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15622" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15623" class="Symbol">)))</a> <a id="15628" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="15631" href="Algebra.Lattice.Properties.BooleanAlgebra.html#5721" class="Function">deMorgan₁</a> <a id="15641" class="Symbol">_</a> <a id="15643" class="Symbol">_</a> <a id="15645" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="15653" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15655" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="15657" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15659" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15661" class="Symbol">((</a><a id="15663" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15665" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15667" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15668" class="Symbol">)</a> <a id="15670" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15672" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15674" class="Symbol">(</a><a id="15675" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15677" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15679" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15680" class="Symbol">))</a> <a id="15684" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="15687" href="Algebra.Lattice.Structures.html#4639" class="Function">∨-congˡ</a> <a id="15695" href="Algebra.Lattice.Properties.BooleanAlgebra.html#15018" class="Function">lem₄′</a> <a id="15701" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="15709" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15711" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="15713" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15715" class="Symbol">((</a><a id="15717" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15719" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15721" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15723" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15724" class="Symbol">)</a> <a id="15726" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="15728" class="Symbol">(</a><a id="15729" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15731" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15733" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15735" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15736" class="Symbol">))</a> <a id="15740" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="15743" href="Algebra.Lattice.Structures.html#4978" class="Function">∨-distribˡ-∧</a> <a id="15756" class="Symbol">_</a> <a id="15758" class="Symbol">_</a> <a id="15760" class="Symbol">_</a> <a id="15762" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="15770" class="Symbol">(</a><a id="15771" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15773" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="15775" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15777" class="Symbol">(</a><a id="15778" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15784" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15786" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15788" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15789" class="Symbol">))</a> <a id="15792" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="15800" class="Symbol">(</a><a id="15801" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15803" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="15805" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15807" class="Symbol">(</a><a id="15808" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15810" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a> <a id="15812" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15814" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15815" class="Symbol">))</a> <a id="15831" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="15834" href="Algebra.Lattice.Structures.html#4147" class="Function">∨-assoc</a> <a id="15842" class="Symbol">_</a> <a id="15844" class="Symbol">_</a> <a id="15846" class="Symbol">_</a> <a id="15848" href="Function.Base.html#4322" class="Function Operator">⟨</a> <a id="15850" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a> <a id="15857" href="Function.Base.html#4322" class="Function Operator">⟩</a> <a id="15859" href="Algebra.Lattice.Structures.html#4147" class="Function">∨-assoc</a> <a id="15867" class="Symbol">_</a> <a id="15869" class="Symbol">_</a> <a id="15871" class="Symbol">_</a> <a id="15873" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="15881" class="Symbol">((</a><a id="15883" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15885" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="15887" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15889" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="15890" class="Symbol">)</a> <a id="15896" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15898" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15900" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15901" class="Symbol">)</a> <a id="15903" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="15911" class="Symbol">((</a><a id="15913" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15915" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="15917" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15919" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15921" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="15922" class="Symbol">)</a> <a id="15924" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15926" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15927" class="Symbol">)</a> <a id="15942" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="15945" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="15952" class="Symbol">_</a> <a id="15954" class="Symbol">_</a> <a id="15956" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="15964" class="Symbol">((</a><a id="15966" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15968" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="15970" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15972" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15974" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="15975" class="Symbol">)</a> <a id="15977" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15979" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="15980" class="Symbol">)</a> <a id="15982" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="15990" class="Symbol">((</a><a id="15992" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="15994" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="15996" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="15998" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="15999" class="Symbol">)</a> <a id="16005" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16007" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16009" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16010" class="Symbol">)</a> <a id="16021" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="16028" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16028" class="Function">lem₅</a> <a id="16033" class="Symbol">=</a> <a id="16035" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
||
<a id="16047" class="Symbol">((</a><a id="16049" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16051" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16053" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16055" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16056" class="Symbol">)</a> <a id="16058" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16060" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16062" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16063" class="Symbol">)</a> <a id="16065" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="16073" class="Symbol">(((</a><a id="16076" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16078" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16080" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16082" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16084" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16085" class="Symbol">)</a> <a id="16087" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16089" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16090" class="Symbol">)</a> <a id="16092" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="16094" class="Symbol">((</a><a id="16096" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16098" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16100" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16102" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16103" class="Symbol">)</a> <a id="16105" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16107" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16109" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16110" class="Symbol">))</a> <a id="16116" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">≈⟨</a> <a id="16119" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="16127" class="Symbol">_</a> <a id="16129" class="Symbol">_</a> <a id="16131" class="Symbol">_</a> <a id="16133" href="Relation.Binary.Reasoning.Syntax.html#7074" class="Function">⟨</a>
|
||
<a id="16141" class="Symbol">(((</a><a id="16144" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16146" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16148" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16150" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16151" class="Symbol">)</a> <a id="16153" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16155" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16157" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16158" class="Symbol">)</a> <a id="16160" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="16162" class="Symbol">((</a><a id="16164" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16166" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16168" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16170" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16172" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16173" class="Symbol">)</a> <a id="16175" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16177" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16178" class="Symbol">))</a> <a id="16181" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="16189" class="Symbol">((</a><a id="16191" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16193" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16195" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16197" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16198" class="Symbol">)</a> <a id="16200" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16202" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16204" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16205" class="Symbol">)</a> <a id="16232" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="16235" href="Algebra.Lattice.Structures.html#4577" class="Function">∧-congʳ</a> <a id="16243" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="16245" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a> <a id="16252" class="Symbol">_</a> <a id="16254" class="Symbol">_</a> <a id="16256" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="16264" class="Symbol">(((</a><a id="16267" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16269" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16271" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16273" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16275" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16276" class="Symbol">)</a> <a id="16278" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16280" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16281" class="Symbol">)</a> <a id="16283" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="16285" class="Symbol">((</a><a id="16287" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16289" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16291" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16293" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16294" class="Symbol">)</a> <a id="16296" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16298" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16300" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16301" class="Symbol">))</a> <a id="16304" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="16312" class="Symbol">((</a><a id="16314" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16316" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16318" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16320" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16321" class="Symbol">)</a> <a id="16323" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16325" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16327" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16328" class="Symbol">)</a> <a id="16355" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="16358" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a> <a id="16366" class="Symbol">_</a> <a id="16368" class="Symbol">_</a> <a id="16370" class="Symbol">_</a> <a id="16372" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
||
<a id="16380" class="Symbol">((</a><a id="16382" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16384" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16386" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16388" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16390" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16391" class="Symbol">)</a> <a id="16393" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16395" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16396" class="Symbol">)</a> <a id="16398" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a>
|
||
<a id="16406" class="Symbol">(((</a><a id="16409" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16411" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16413" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16415" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16416" class="Symbol">)</a> <a id="16418" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16420" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16422" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16423" class="Symbol">)</a> <a id="16425" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="16427" class="Symbol">((</a><a id="16429" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16431" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12720" class="Bound">x</a> <a id="16433" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16435" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12722" class="Bound">y</a><a id="16436" class="Symbol">)</a> <a id="16438" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="16440" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="16442" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12724" class="Bound">z</a><a id="16443" class="Symbol">))</a> <a id="16449" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
||
|
||
<a id="XorRing.⊕-isMagma"></a><a id="16454" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16454" class="Function">⊕-isMagma</a> <a id="16464" class="Symbol">:</a> <a id="16466" href="Algebra.Structures.html#1225" class="Record">IsMagma</a> <a id="16474" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="16480" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16454" class="Function">⊕-isMagma</a> <a id="16490" class="Symbol">=</a> <a id="16492" class="Keyword">record</a>
|
||
<a id="16503" class="Symbol">{</a> <a id="16505" href="Algebra.Structures.html#1277" class="Field">isEquivalence</a> <a id="16519" class="Symbol">=</a> <a id="16521" href="Algebra.Lattice.Structures.html#4075" class="Function">isEquivalence</a>
|
||
<a id="16539" class="Symbol">;</a> <a id="16541" href="Algebra.Structures.html#1315" class="Field">∙-cong</a> <a id="16555" class="Symbol">=</a> <a id="16557" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7595" class="Function">⊕-cong</a>
|
||
<a id="16568" class="Symbol">}</a>
|
||
|
||
<a id="XorRing.⊕-isSemigroup"></a><a id="16573" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16573" class="Function">⊕-isSemigroup</a> <a id="16587" class="Symbol">:</a> <a id="16589" href="Algebra.Structures.html#2897" class="Record">IsSemigroup</a> <a id="16601" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a>
|
||
<a id="16607" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16573" class="Function">⊕-isSemigroup</a> <a id="16621" class="Symbol">=</a> <a id="16623" class="Keyword">record</a>
|
||
<a id="16634" class="Symbol">{</a> <a id="16636" href="Algebra.Structures.html#2953" class="Field">isMagma</a> <a id="16644" class="Symbol">=</a> <a id="16646" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16454" class="Function">⊕-isMagma</a>
|
||
<a id="16660" class="Symbol">;</a> <a id="16662" href="Algebra.Structures.html#2977" class="Field">assoc</a> <a id="16670" class="Symbol">=</a> <a id="16672" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12684" class="Function">⊕-assoc</a>
|
||
<a id="16684" class="Symbol">}</a>
|
||
|
||
<a id="XorRing.⊕-⊥-isMonoid"></a><a id="16689" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16689" class="Function">⊕-⊥-isMonoid</a> <a id="16702" class="Symbol">:</a> <a id="16704" href="Algebra.Structures.html#3974" class="Record">IsMonoid</a> <a id="16713" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a> <a id="16717" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a>
|
||
<a id="16721" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16689" class="Function">⊕-⊥-isMonoid</a> <a id="16734" class="Symbol">=</a> <a id="16736" class="Keyword">record</a>
|
||
<a id="16747" class="Symbol">{</a> <a id="16749" href="Algebra.Structures.html#4035" class="Field">isSemigroup</a> <a id="16761" class="Symbol">=</a> <a id="16763" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16573" class="Function">⊕-isSemigroup</a>
|
||
<a id="16781" class="Symbol">;</a> <a id="16783" href="Algebra.Structures.html#4067" class="Field">identity</a> <a id="16795" class="Symbol">=</a> <a id="16797" href="Algebra.Lattice.Properties.BooleanAlgebra.html#9947" class="Function">⊕-identity</a>
|
||
<a id="16812" class="Symbol">}</a>
|
||
|
||
<a id="XorRing.⊕-⊥-isGroup"></a><a id="16817" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16817" class="Function">⊕-⊥-isGroup</a> <a id="16829" class="Symbol">:</a> <a id="16831" href="Algebra.Structures.html#6188" class="Record">IsGroup</a> <a id="16839" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a> <a id="16843" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="16845" href="Function.Base.html#704" class="Function">id</a>
|
||
<a id="16850" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16817" class="Function">⊕-⊥-isGroup</a> <a id="16862" class="Symbol">=</a> <a id="16864" class="Keyword">record</a>
|
||
<a id="16875" class="Symbol">{</a> <a id="16877" href="Algebra.Structures.html#6264" class="Field">isMonoid</a> <a id="16886" class="Symbol">=</a> <a id="16888" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16689" class="Function">⊕-⊥-isMonoid</a>
|
||
<a id="16905" class="Symbol">;</a> <a id="16907" href="Algebra.Structures.html#6295" class="Field">inverse</a> <a id="16916" class="Symbol">=</a> <a id="16918" href="Algebra.Lattice.Properties.BooleanAlgebra.html#10336" class="Function">⊕-inverse</a>
|
||
<a id="16932" class="Symbol">;</a> <a id="16934" href="Algebra.Structures.html#6329" class="Field">⁻¹-cong</a> <a id="16943" class="Symbol">=</a> <a id="16945" href="Function.Base.html#704" class="Function">id</a>
|
||
<a id="16952" class="Symbol">}</a>
|
||
|
||
<a id="XorRing.⊕-⊥-isAbelianGroup"></a><a id="16957" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16957" class="Function">⊕-⊥-isAbelianGroup</a> <a id="16976" class="Symbol">:</a> <a id="16978" href="Algebra.Structures.html#7222" class="Record">IsAbelianGroup</a> <a id="16993" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a> <a id="16997" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="16999" href="Function.Base.html#704" class="Function">id</a>
|
||
<a id="17004" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16957" class="Function">⊕-⊥-isAbelianGroup</a> <a id="17023" class="Symbol">=</a> <a id="17025" class="Keyword">record</a>
|
||
<a id="17036" class="Symbol">{</a> <a id="17038" href="Algebra.Structures.html#7324" class="Field">isGroup</a> <a id="17046" class="Symbol">=</a> <a id="17048" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16817" class="Function">⊕-⊥-isGroup</a>
|
||
<a id="17064" class="Symbol">;</a> <a id="17066" href="Algebra.Structures.html#7353" class="Field">comm</a> <a id="17074" class="Symbol">=</a> <a id="17076" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7887" class="Function">⊕-comm</a>
|
||
<a id="17087" class="Symbol">}</a>
|
||
|
||
<a id="XorRing.⊕-∧-isRing"></a><a id="17092" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17092" class="Function">⊕-∧-isRing</a> <a id="17103" class="Symbol">:</a> <a id="17105" href="Algebra.Structures.html#21839" class="Record">IsRing</a> <a id="17112" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a> <a id="17116" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="17120" href="Function.Base.html#704" class="Function">id</a> <a id="17123" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="17125" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a>
|
||
<a id="17129" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17092" class="Function">⊕-∧-isRing</a> <a id="17140" class="Symbol">=</a> <a id="17142" class="Keyword">record</a>
|
||
<a id="17153" class="Symbol">{</a> <a id="17155" href="Algebra.Structures.html#21917" class="Field">+-isAbelianGroup</a> <a id="17172" class="Symbol">=</a> <a id="17174" href="Algebra.Lattice.Properties.BooleanAlgebra.html#16957" class="Function">⊕-⊥-isAbelianGroup</a>
|
||
<a id="17197" class="Symbol">;</a> <a id="17199" href="Algebra.Structures.html#21963" class="Field">*-cong</a> <a id="17206" class="Symbol">=</a> <a id="17208" href="Algebra.Lattice.Structures.html#4282" class="Function">∧-cong</a>
|
||
<a id="17219" class="Symbol">;</a> <a id="17221" href="Algebra.Structures.html#21999" class="Field">*-assoc</a> <a id="17229" class="Symbol">=</a> <a id="17231" href="Algebra.Lattice.Structures.html#4248" class="Function">∧-assoc</a>
|
||
<a id="17243" class="Symbol">;</a> <a id="17245" href="Algebra.Structures.html#22036" class="Field">*-identity</a> <a id="17256" class="Symbol">=</a> <a id="17258" href="Algebra.Lattice.Properties.BooleanAlgebra.html#1994" class="Function">∧-identity</a>
|
||
<a id="17273" class="Symbol">;</a> <a id="17275" href="Algebra.Structures.html#22073" class="Field">distrib</a> <a id="17283" class="Symbol">=</a> <a id="17285" href="Algebra.Lattice.Properties.BooleanAlgebra.html#12121" class="Function">∧-distrib-⊕</a>
|
||
<a id="17301" class="Symbol">}</a>
|
||
|
||
<a id="XorRing.⊕-∧-isCommutativeRing"></a><a id="17306" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17306" class="Function">⊕-∧-isCommutativeRing</a> <a id="17328" class="Symbol">:</a> <a id="17330" href="Algebra.Structures.html#23275" class="Record">IsCommutativeRing</a> <a id="17348" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7451" class="Function Operator">_⊕_</a> <a id="17352" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">_∧_</a> <a id="17356" href="Function.Base.html#704" class="Function">id</a> <a id="17359" href="Algebra.Lattice.Bundles.html#5872" class="Field">⊥</a> <a id="17361" href="Algebra.Lattice.Bundles.html#5841" class="Field">⊤</a>
|
||
<a id="17365" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17306" class="Function">⊕-∧-isCommutativeRing</a> <a id="17387" class="Symbol">=</a> <a id="17389" class="Keyword">record</a>
|
||
<a id="17400" class="Symbol">{</a> <a id="17402" href="Algebra.Structures.html#23372" class="Field">isRing</a> <a id="17409" class="Symbol">=</a> <a id="17411" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17092" class="Function">⊕-∧-isRing</a>
|
||
<a id="17426" class="Symbol">;</a> <a id="17428" href="Algebra.Structures.html#23404" class="Field">*-comm</a> <a id="17435" class="Symbol">=</a> <a id="17437" href="Algebra.Lattice.Structures.html#4214" class="Function">∧-comm</a>
|
||
<a id="17448" class="Symbol">}</a>
|
||
|
||
<a id="XorRing.⊕-∧-commutativeRing"></a><a id="17453" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17453" class="Function">⊕-∧-commutativeRing</a> <a id="17473" class="Symbol">:</a> <a id="17475" href="Algebra.Bundles.html#27248" class="Record">CommutativeRing</a> <a id="17491" class="Symbol">_</a> <a id="17493" class="Symbol">_</a>
|
||
<a id="17497" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17453" class="Function">⊕-∧-commutativeRing</a> <a id="17517" class="Symbol">=</a> <a id="17519" class="Keyword">record</a>
|
||
<a id="17530" class="Symbol">{</a> <a id="17532" href="Algebra.Bundles.html#27605" class="Field">isCommutativeRing</a> <a id="17550" class="Symbol">=</a> <a id="17552" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17306" class="Function">⊕-∧-isCommutativeRing</a>
|
||
<a id="17578" class="Symbol">}</a>
|
||
|
||
|
||
<a id="17582" class="Keyword">infixl</a> <a id="17589" class="Number">6</a> <a id="17591" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17596" class="Function Operator">_⊕_</a>
|
||
|
||
<a id="_⊕_"></a><a id="17596" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17596" class="Function Operator">_⊕_</a> <a id="17600" class="Symbol">:</a> <a id="17602" href="Algebra.Core.html#527" class="Function">Op₂</a> <a id="17606" href="Algebra.Lattice.Bundles.html#5670" class="Field">Carrier</a>
|
||
<a id="17614" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17614" class="Bound">x</a> <a id="17616" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17596" class="Function Operator">⊕</a> <a id="17618" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17618" class="Bound">y</a> <a id="17620" class="Symbol">=</a> <a id="17622" class="Symbol">(</a><a id="17623" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17614" class="Bound">x</a> <a id="17625" href="Algebra.Lattice.Bundles.html#5736" class="Field Operator">∨</a> <a id="17627" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17618" class="Bound">y</a><a id="17628" class="Symbol">)</a> <a id="17630" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="17632" href="Algebra.Lattice.Bundles.html#5806" class="Field Operator">¬</a> <a id="17634" class="Symbol">(</a><a id="17635" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17614" class="Bound">x</a> <a id="17637" href="Algebra.Lattice.Bundles.html#5771" class="Field Operator">∧</a> <a id="17639" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17618" class="Bound">y</a><a id="17640" class="Symbol">)</a>
|
||
|
||
<a id="17643" class="Keyword">module</a> <a id="DefaultXorRing"></a><a id="17650" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17650" class="Module">DefaultXorRing</a> <a id="17665" class="Symbol">=</a> <a id="17667" href="Algebra.Lattice.Properties.BooleanAlgebra.html#7330" class="Module">XorRing</a> <a id="17675" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17596" class="Function Operator">_⊕_</a> <a id="17679" class="Symbol">(λ</a> <a id="17682" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17682" class="Bound">_</a> <a id="17684" href="Algebra.Lattice.Properties.BooleanAlgebra.html#17684" class="Bound">_</a> <a id="17686" class="Symbol">→</a> <a id="17688" href="Relation.Binary.Structures.html#1596" class="Function">refl</a><a id="17692" class="Symbol">)</a>
|
||
</pre></body></html> |