mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
200 lines
No EOL
56 KiB
HTML
200 lines
No EOL
56 KiB
HTML
<!DOCTYPE HTML>
|
||
<html><head><meta charset="utf-8"><title>Relation.Nullary.Decidable.Core</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">-- Operations on and properties of decidable relations</a>
|
||
<a id="161" class="Comment">--</a>
|
||
<a id="164" class="Comment">-- This file contains some core definitions which are re-exported by</a>
|
||
<a id="233" class="Comment">-- Relation.Nullary.Decidable</a>
|
||
<a id="263" class="Comment">------------------------------------------------------------------------</a>
|
||
|
||
<a id="337" class="Symbol">{-#</a> <a id="341" class="Keyword">OPTIONS</a> <a id="349" class="Pragma">--cubical-compatible</a> <a id="370" class="Pragma">--safe</a> <a id="377" class="Symbol">#-}</a>
|
||
|
||
<a id="382" class="Keyword">module</a> <a id="389" href="Relation.Nullary.Decidable.Core.html" class="Module">Relation.Nullary.Decidable.Core</a> <a id="421" class="Keyword">where</a>
|
||
|
||
<a id="428" class="Keyword">open</a> <a id="433" class="Keyword">import</a> <a id="440" href="Level.html" class="Module">Level</a> <a id="446" class="Keyword">using</a> <a id="452" class="Symbol">(</a><a id="453" href="Agda.Primitive.html#742" class="Postulate">Level</a><a id="458" class="Symbol">;</a> <a id="460" href="Level.html#409" class="Record">Lift</a><a id="464" class="Symbol">)</a>
|
||
<a id="466" class="Keyword">open</a> <a id="471" class="Keyword">import</a> <a id="478" href="Data.Bool.Base.html" class="Module">Data.Bool.Base</a> <a id="493" class="Keyword">using</a> <a id="499" class="Symbol">(</a><a id="500" href="Agda.Builtin.Bool.html#173" class="Datatype">Bool</a><a id="504" class="Symbol">;</a> <a id="506" href="Data.Bool.Base.html#1348" class="Function">T</a><a id="507" class="Symbol">;</a> <a id="509" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a><a id="514" class="Symbol">;</a> <a id="516" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a><a id="520" class="Symbol">;</a> <a id="522" href="Data.Bool.Base.html#941" class="Function">not</a><a id="525" class="Symbol">;</a> <a id="527" href="Data.Bool.Base.html#995" class="Function Operator">_∧_</a><a id="530" class="Symbol">;</a> <a id="532" href="Data.Bool.Base.html#1053" class="Function Operator">_∨_</a><a id="535" class="Symbol">)</a>
|
||
<a id="537" class="Keyword">open</a> <a id="542" class="Keyword">import</a> <a id="549" href="Data.Unit.Base.html" class="Module">Data.Unit.Base</a> <a id="564" class="Keyword">using</a> <a id="570" class="Symbol">(</a><a id="571" href="Agda.Builtin.Unit.html#175" class="Record">⊤</a><a id="572" class="Symbol">)</a>
|
||
<a id="574" class="Keyword">open</a> <a id="579" class="Keyword">import</a> <a id="586" href="Data.Empty.html" class="Module">Data.Empty</a> <a id="597" class="Keyword">using</a> <a id="603" class="Symbol">(</a><a id="604" href="Data.Empty.html#895" class="Function">⊥</a><a id="605" class="Symbol">)</a>
|
||
<a id="607" class="Keyword">open</a> <a id="612" class="Keyword">import</a> <a id="619" href="Data.Empty.Irrelevant.html" class="Module">Data.Empty.Irrelevant</a> <a id="641" class="Keyword">using</a> <a id="647" class="Symbol">(</a><a id="648" href="Data.Empty.Irrelevant.html#336" class="Function">⊥-elim</a><a id="654" class="Symbol">)</a>
|
||
<a id="656" class="Keyword">open</a> <a id="661" class="Keyword">import</a> <a id="668" href="Data.Product.Base.html" class="Module">Data.Product.Base</a> <a id="686" class="Keyword">using</a> <a id="692" class="Symbol">(</a><a id="693" href="Data.Product.Base.html#1618" class="Function Operator">_×_</a><a id="696" class="Symbol">)</a>
|
||
<a id="698" class="Keyword">open</a> <a id="703" class="Keyword">import</a> <a id="710" href="Data.Sum.Base.html" class="Module">Data.Sum.Base</a> <a id="724" class="Keyword">using</a> <a id="730" class="Symbol">(</a><a id="731" href="Data.Sum.Base.html#625" class="Datatype Operator">_⊎_</a><a id="734" class="Symbol">)</a>
|
||
<a id="736" class="Keyword">open</a> <a id="741" class="Keyword">import</a> <a id="748" href="Function.Base.html" class="Module">Function.Base</a> <a id="762" class="Keyword">using</a> <a id="768" class="Symbol">(</a><a id="769" href="Function.Base.html#1115" class="Function Operator">_∘_</a><a id="772" class="Symbol">;</a> <a id="774" href="Function.Base.html#725" class="Function">const</a><a id="779" class="Symbol">;</a> <a id="781" href="Function.Base.html#1974" class="Function Operator">_$_</a><a id="784" class="Symbol">;</a> <a id="786" href="Function.Base.html#1638" class="Function">flip</a><a id="790" class="Symbol">)</a>
|
||
<a id="792" class="Keyword">open</a> <a id="797" class="Keyword">import</a> <a id="804" href="Relation.Nullary.Reflects.html" class="Module">Relation.Nullary.Reflects</a>
|
||
<a id="830" class="Keyword">open</a> <a id="835" class="Keyword">import</a> <a id="842" href="Relation.Nullary.Negation.Core.html" class="Module">Relation.Nullary.Negation.Core</a>
|
||
|
||
<a id="874" class="Keyword">private</a>
|
||
<a id="884" class="Keyword">variable</a>
|
||
<a id="897" href="Relation.Nullary.Decidable.Core.html#897" class="Generalizable">a</a> <a id="899" href="Relation.Nullary.Decidable.Core.html#899" class="Generalizable">b</a> <a id="901" class="Symbol">:</a> <a id="903" href="Agda.Primitive.html#742" class="Postulate">Level</a>
|
||
<a id="913" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="915" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a> <a id="917" class="Symbol">:</a> <a id="919" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="923" href="Relation.Nullary.Decidable.Core.html#897" class="Generalizable">a</a>
|
||
|
||
<a id="926" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="999" class="Comment">-- Definition.</a>
|
||
|
||
<a id="1015" class="Comment">-- Decidability proofs have two parts: the `does` term which contains</a>
|
||
<a id="1085" class="Comment">-- the boolean result and the `proof` term which contains a proof that</a>
|
||
<a id="1156" class="Comment">-- reflects the boolean result. This definition allows the boolean</a>
|
||
<a id="1223" class="Comment">-- part of the decision procedure to compute independently from the</a>
|
||
<a id="1291" class="Comment">-- proof. This leads to better computational behaviour when we only care</a>
|
||
<a id="1364" class="Comment">-- about the result and not the proof. See README.Design.Decidability</a>
|
||
<a id="1434" class="Comment">-- for further details.</a>
|
||
|
||
<a id="1459" class="Keyword">infix</a> <a id="1465" class="Number">2</a> <a id="1467" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">_because_</a>
|
||
|
||
<a id="1478" class="Keyword">record</a> <a id="Dec"></a><a id="1485" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="1489" class="Symbol">(</a><a id="1490" href="Relation.Nullary.Decidable.Core.html#1490" class="Bound">A</a> <a id="1492" class="Symbol">:</a> <a id="1494" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1498" href="Relation.Nullary.Decidable.Core.html#897" class="Generalizable">a</a><a id="1499" class="Symbol">)</a> <a id="1501" class="Symbol">:</a> <a id="1503" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1507" href="Relation.Nullary.Decidable.Core.html#1498" class="Bound">a</a> <a id="1509" class="Keyword">where</a>
|
||
<a id="1517" class="Keyword">constructor</a> <a id="_because_"></a><a id="1529" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">_because_</a>
|
||
<a id="1541" class="Keyword">field</a>
|
||
<a id="Dec.does"></a><a id="1551" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="1557" class="Symbol">:</a> <a id="1559" href="Agda.Builtin.Bool.html#173" class="Datatype">Bool</a>
|
||
<a id="Dec.proof"></a><a id="1568" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="1574" class="Symbol">:</a> <a id="1576" href="Relation.Nullary.Reflects.html#938" class="Datatype">Reflects</a> <a id="1585" href="Relation.Nullary.Decidable.Core.html#1490" class="Bound">A</a> <a id="1587" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a>
|
||
|
||
<a id="1593" class="Keyword">open</a> <a id="1598" href="Relation.Nullary.Decidable.Core.html#1485" class="Module">Dec</a> <a id="1602" class="Keyword">public</a>
|
||
|
||
<a id="1610" class="Keyword">pattern</a> <a id="yes"></a><a id="1618" href="Relation.Nullary.Decidable.Core.html#1618" class="InductiveConstructor">yes</a> <a id="1622" href="Relation.Nullary.Decidable.Core.html#1645" class="Bound">a</a> <a id="1624" class="Symbol">=</a> <a id="1627" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="1632" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="1640" href="Relation.Nullary.Reflects.html#982" class="InductiveConstructor">ofʸ</a> <a id="1645" href="Relation.Nullary.Decidable.Core.html#1645" class="Bound">a</a>
|
||
<a id="1647" class="Keyword">pattern</a> <a id="no"></a><a id="1655" href="Relation.Nullary.Decidable.Core.html#1655" class="InductiveConstructor">no</a> <a id="1658" href="Relation.Nullary.Decidable.Core.html#1681" class="Bound">¬a</a> <a id="1661" class="Symbol">=</a> <a id="1663" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="1669" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="1677" href="Relation.Nullary.Reflects.html#1019" class="InductiveConstructor">ofⁿ</a> <a id="1681" href="Relation.Nullary.Decidable.Core.html#1681" class="Bound">¬a</a>
|
||
|
||
<a id="1685" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="1758" class="Comment">-- Flattening</a>
|
||
|
||
<a id="1773" class="Keyword">module</a> <a id="1780" href="Relation.Nullary.Decidable.Core.html#1780" class="Module">_</a> <a id="1782" class="Symbol">{</a><a id="1783" href="Relation.Nullary.Decidable.Core.html#1783" class="Bound">A</a> <a id="1785" class="Symbol">:</a> <a id="1787" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1791" href="Relation.Nullary.Decidable.Core.html#897" class="Generalizable">a</a><a id="1792" class="Symbol">}</a> <a id="1794" class="Keyword">where</a>
|
||
|
||
<a id="1803" href="Relation.Nullary.Decidable.Core.html#1803" class="Function">From-yes</a> <a id="1812" class="Symbol">:</a> <a id="1814" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="1818" href="Relation.Nullary.Decidable.Core.html#1783" class="Bound">A</a> <a id="1820" class="Symbol">→</a> <a id="1822" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1826" href="Relation.Nullary.Decidable.Core.html#1791" class="Bound">a</a>
|
||
<a id="1830" href="Relation.Nullary.Decidable.Core.html#1803" class="Function">From-yes</a> <a id="1839" class="Symbol">(</a><a id="1840" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="1846" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="1854" class="Symbol">_)</a> <a id="1857" class="Symbol">=</a> <a id="1859" href="Relation.Nullary.Decidable.Core.html#1783" class="Bound">A</a>
|
||
<a id="1863" href="Relation.Nullary.Decidable.Core.html#1803" class="Function">From-yes</a> <a id="1872" class="Symbol">(</a><a id="1873" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="1879" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="1887" class="Symbol">_)</a> <a id="1890" class="Symbol">=</a> <a id="1892" href="Level.html#409" class="Record">Lift</a> <a id="1897" href="Relation.Nullary.Decidable.Core.html#1791" class="Bound">a</a> <a id="1899" href="Agda.Builtin.Unit.html#175" class="Record">⊤</a>
|
||
|
||
<a id="1904" href="Relation.Nullary.Decidable.Core.html#1904" class="Function">From-no</a> <a id="1912" class="Symbol">:</a> <a id="1914" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="1918" href="Relation.Nullary.Decidable.Core.html#1783" class="Bound">A</a> <a id="1920" class="Symbol">→</a> <a id="1922" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1926" href="Relation.Nullary.Decidable.Core.html#1791" class="Bound">a</a>
|
||
<a id="1930" href="Relation.Nullary.Decidable.Core.html#1904" class="Function">From-no</a> <a id="1938" class="Symbol">(</a><a id="1939" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="1945" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="1953" class="Symbol">_)</a> <a id="1956" class="Symbol">=</a> <a id="1958" href="Relation.Nullary.Negation.Core.html#698" class="Function Operator">¬</a> <a id="1960" href="Relation.Nullary.Decidable.Core.html#1783" class="Bound">A</a>
|
||
<a id="1964" href="Relation.Nullary.Decidable.Core.html#1904" class="Function">From-no</a> <a id="1972" class="Symbol">(</a><a id="1973" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="1979" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="1987" class="Symbol">_)</a> <a id="1990" class="Symbol">=</a> <a id="1992" href="Level.html#409" class="Record">Lift</a> <a id="1997" href="Relation.Nullary.Decidable.Core.html#1791" class="Bound">a</a> <a id="1999" href="Agda.Builtin.Unit.html#175" class="Record">⊤</a>
|
||
|
||
<a id="2002" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="2075" class="Comment">-- Recompute</a>
|
||
|
||
<a id="2089" class="Comment">-- Given an irrelevant proof of a decidable type, a proof can</a>
|
||
<a id="2151" class="Comment">-- be recomputed and subsequently used in relevant contexts.</a>
|
||
<a id="recompute"></a><a id="2212" href="Relation.Nullary.Decidable.Core.html#2212" class="Function">recompute</a> <a id="2222" class="Symbol">:</a> <a id="2224" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2228" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="2230" class="Symbol">→</a> <a id="2232" class="Symbol">.</a><a id="2233" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="2235" class="Symbol">→</a> <a id="2237" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a>
|
||
<a id="2239" href="Relation.Nullary.Decidable.Core.html#2212" class="Function">recompute</a> <a id="2249" class="Symbol">(</a><a id="2250" href="Relation.Nullary.Decidable.Core.html#1618" class="InductiveConstructor">yes</a> <a id="2254" href="Relation.Nullary.Decidable.Core.html#2254" class="Bound">a</a><a id="2255" class="Symbol">)</a> <a id="2257" class="Symbol">_</a> <a id="2259" class="Symbol">=</a> <a id="2261" href="Relation.Nullary.Decidable.Core.html#2254" class="Bound">a</a>
|
||
<a id="2263" href="Relation.Nullary.Decidable.Core.html#2212" class="Function">recompute</a> <a id="2273" class="Symbol">(</a><a id="2274" href="Relation.Nullary.Decidable.Core.html#1655" class="InductiveConstructor">no</a> <a id="2277" href="Relation.Nullary.Decidable.Core.html#2277" class="Bound">¬a</a><a id="2279" class="Symbol">)</a> <a id="2281" href="Relation.Nullary.Decidable.Core.html#2281" class="Bound">a</a> <a id="2283" class="Symbol">=</a> <a id="2285" href="Data.Empty.Irrelevant.html#336" class="Function">⊥-elim</a> <a id="2292" class="Symbol">(</a><a id="2293" href="Relation.Nullary.Decidable.Core.html#2277" class="Bound">¬a</a> <a id="2296" href="Relation.Nullary.Decidable.Core.html#2281" class="Bound">a</a><a id="2297" class="Symbol">)</a>
|
||
|
||
<a id="2300" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="2373" class="Comment">-- Interaction with negation, sum, product etc.</a>
|
||
|
||
<a id="2422" class="Keyword">infixr</a> <a id="2429" class="Number">1</a> <a id="2431" href="Relation.Nullary.Decidable.Core.html#2739" class="Function Operator">_⊎-dec_</a>
|
||
<a id="2439" class="Keyword">infixr</a> <a id="2446" class="Number">2</a> <a id="2448" href="Relation.Nullary.Decidable.Core.html#2609" class="Function Operator">_×-dec_</a> <a id="2456" href="Relation.Nullary.Decidable.Core.html#2869" class="Function Operator">_→-dec_</a>
|
||
|
||
<a id="T?"></a><a id="2465" href="Relation.Nullary.Decidable.Core.html#2465" class="Function">T?</a> <a id="2468" class="Symbol">:</a> <a id="2470" class="Symbol">∀</a> <a id="2472" href="Relation.Nullary.Decidable.Core.html#2472" class="Bound">x</a> <a id="2474" class="Symbol">→</a> <a id="2476" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2480" class="Symbol">(</a><a id="2481" href="Data.Bool.Base.html#1348" class="Function">T</a> <a id="2483" href="Relation.Nullary.Decidable.Core.html#2472" class="Bound">x</a><a id="2484" class="Symbol">)</a>
|
||
<a id="2486" href="Relation.Nullary.Decidable.Core.html#2465" class="Function">T?</a> <a id="2489" href="Relation.Nullary.Decidable.Core.html#2489" class="Bound">x</a> <a id="2491" class="Symbol">=</a> <a id="2493" href="Relation.Nullary.Decidable.Core.html#2489" class="Bound">x</a> <a id="2495" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="2503" href="Relation.Nullary.Reflects.html#1858" class="Function">T-reflects</a> <a id="2514" href="Relation.Nullary.Decidable.Core.html#2489" class="Bound">x</a>
|
||
|
||
<a id="¬?"></a><a id="2517" href="Relation.Nullary.Decidable.Core.html#2517" class="Function">¬?</a> <a id="2520" class="Symbol">:</a> <a id="2522" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2526" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="2528" class="Symbol">→</a> <a id="2530" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2534" class="Symbol">(</a><a id="2535" href="Relation.Nullary.Negation.Core.html#698" class="Function Operator">¬</a> <a id="2537" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="2538" class="Symbol">)</a>
|
||
<a id="2540" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2546" class="Symbol">(</a><a id="2547" href="Relation.Nullary.Decidable.Core.html#2517" class="Function">¬?</a> <a id="2550" href="Relation.Nullary.Decidable.Core.html#2550" class="Bound">a?</a><a id="2552" class="Symbol">)</a> <a id="2554" class="Symbol">=</a> <a id="2556" href="Data.Bool.Base.html#941" class="Function">not</a> <a id="2560" class="Symbol">(</a><a id="2561" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2566" href="Relation.Nullary.Decidable.Core.html#2550" class="Bound">a?</a><a id="2568" class="Symbol">)</a>
|
||
<a id="2570" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2576" class="Symbol">(</a><a id="2577" href="Relation.Nullary.Decidable.Core.html#2517" class="Function">¬?</a> <a id="2580" href="Relation.Nullary.Decidable.Core.html#2580" class="Bound">a?</a><a id="2582" class="Symbol">)</a> <a id="2584" class="Symbol">=</a> <a id="2586" href="Relation.Nullary.Reflects.html#2000" class="Function">¬-reflects</a> <a id="2597" class="Symbol">(</a><a id="2598" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2604" href="Relation.Nullary.Decidable.Core.html#2580" class="Bound">a?</a><a id="2606" class="Symbol">)</a>
|
||
|
||
<a id="_×-dec_"></a><a id="2609" href="Relation.Nullary.Decidable.Core.html#2609" class="Function Operator">_×-dec_</a> <a id="2617" class="Symbol">:</a> <a id="2619" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2623" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="2625" class="Symbol">→</a> <a id="2627" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2631" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a> <a id="2633" class="Symbol">→</a> <a id="2635" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2639" class="Symbol">(</a><a id="2640" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="2642" href="Data.Product.Base.html#1618" class="Function Operator">×</a> <a id="2644" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a><a id="2645" class="Symbol">)</a>
|
||
<a id="2647" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2653" class="Symbol">(</a><a id="2654" href="Relation.Nullary.Decidable.Core.html#2654" class="Bound">a?</a> <a id="2657" href="Relation.Nullary.Decidable.Core.html#2609" class="Function Operator">×-dec</a> <a id="2663" href="Relation.Nullary.Decidable.Core.html#2663" class="Bound">b?</a><a id="2665" class="Symbol">)</a> <a id="2667" class="Symbol">=</a> <a id="2669" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2674" href="Relation.Nullary.Decidable.Core.html#2654" class="Bound">a?</a> <a id="2677" href="Data.Bool.Base.html#995" class="Function Operator">∧</a> <a id="2679" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2684" href="Relation.Nullary.Decidable.Core.html#2663" class="Bound">b?</a>
|
||
<a id="2687" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2693" class="Symbol">(</a><a id="2694" href="Relation.Nullary.Decidable.Core.html#2694" class="Bound">a?</a> <a id="2697" href="Relation.Nullary.Decidable.Core.html#2609" class="Function Operator">×-dec</a> <a id="2703" href="Relation.Nullary.Decidable.Core.html#2703" class="Bound">b?</a><a id="2705" class="Symbol">)</a> <a id="2707" class="Symbol">=</a> <a id="2709" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2715" href="Relation.Nullary.Decidable.Core.html#2694" class="Bound">a?</a> <a id="2718" href="Relation.Nullary.Reflects.html#2183" class="Function Operator">×-reflects</a> <a id="2729" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2735" href="Relation.Nullary.Decidable.Core.html#2703" class="Bound">b?</a>
|
||
|
||
<a id="_⊎-dec_"></a><a id="2739" href="Relation.Nullary.Decidable.Core.html#2739" class="Function Operator">_⊎-dec_</a> <a id="2747" class="Symbol">:</a> <a id="2749" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2753" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="2755" class="Symbol">→</a> <a id="2757" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2761" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a> <a id="2763" class="Symbol">→</a> <a id="2765" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2769" class="Symbol">(</a><a id="2770" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="2772" href="Data.Sum.Base.html#625" class="Datatype Operator">⊎</a> <a id="2774" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a><a id="2775" class="Symbol">)</a>
|
||
<a id="2777" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2783" class="Symbol">(</a><a id="2784" href="Relation.Nullary.Decidable.Core.html#2784" class="Bound">a?</a> <a id="2787" href="Relation.Nullary.Decidable.Core.html#2739" class="Function Operator">⊎-dec</a> <a id="2793" href="Relation.Nullary.Decidable.Core.html#2793" class="Bound">b?</a><a id="2795" class="Symbol">)</a> <a id="2797" class="Symbol">=</a> <a id="2799" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2804" href="Relation.Nullary.Decidable.Core.html#2784" class="Bound">a?</a> <a id="2807" href="Data.Bool.Base.html#1053" class="Function Operator">∨</a> <a id="2809" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2814" href="Relation.Nullary.Decidable.Core.html#2793" class="Bound">b?</a>
|
||
<a id="2817" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2823" class="Symbol">(</a><a id="2824" href="Relation.Nullary.Decidable.Core.html#2824" class="Bound">a?</a> <a id="2827" href="Relation.Nullary.Decidable.Core.html#2739" class="Function Operator">⊎-dec</a> <a id="2833" href="Relation.Nullary.Decidable.Core.html#2833" class="Bound">b?</a><a id="2835" class="Symbol">)</a> <a id="2837" class="Symbol">=</a> <a id="2839" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2845" href="Relation.Nullary.Decidable.Core.html#2824" class="Bound">a?</a> <a id="2848" href="Relation.Nullary.Reflects.html#2406" class="Function Operator">⊎-reflects</a> <a id="2859" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2865" href="Relation.Nullary.Decidable.Core.html#2833" class="Bound">b?</a>
|
||
|
||
<a id="_→-dec_"></a><a id="2869" href="Relation.Nullary.Decidable.Core.html#2869" class="Function Operator">_→-dec_</a> <a id="2877" class="Symbol">:</a> <a id="2879" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2883" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="2885" class="Symbol">→</a> <a id="2887" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2891" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a> <a id="2893" class="Symbol">→</a> <a id="2895" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="2899" class="Symbol">(</a><a id="2900" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="2902" class="Symbol">→</a> <a id="2904" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a><a id="2905" class="Symbol">)</a>
|
||
<a id="2907" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2913" class="Symbol">(</a><a id="2914" href="Relation.Nullary.Decidable.Core.html#2914" class="Bound">a?</a> <a id="2917" href="Relation.Nullary.Decidable.Core.html#2869" class="Function Operator">→-dec</a> <a id="2923" href="Relation.Nullary.Decidable.Core.html#2923" class="Bound">b?</a><a id="2925" class="Symbol">)</a> <a id="2927" class="Symbol">=</a> <a id="2929" href="Data.Bool.Base.html#941" class="Function">not</a> <a id="2933" class="Symbol">(</a><a id="2934" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2939" href="Relation.Nullary.Decidable.Core.html#2914" class="Bound">a?</a><a id="2941" class="Symbol">)</a> <a id="2943" href="Data.Bool.Base.html#1053" class="Function Operator">∨</a> <a id="2945" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="2950" href="Relation.Nullary.Decidable.Core.html#2923" class="Bound">b?</a>
|
||
<a id="2953" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2959" class="Symbol">(</a><a id="2960" href="Relation.Nullary.Decidable.Core.html#2960" class="Bound">a?</a> <a id="2963" href="Relation.Nullary.Decidable.Core.html#2869" class="Function Operator">→-dec</a> <a id="2969" href="Relation.Nullary.Decidable.Core.html#2969" class="Bound">b?</a><a id="2971" class="Symbol">)</a> <a id="2973" class="Symbol">=</a> <a id="2975" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="2981" href="Relation.Nullary.Decidable.Core.html#2960" class="Bound">a?</a> <a id="2984" href="Relation.Nullary.Reflects.html#2625" class="Function Operator">→-reflects</a> <a id="2995" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="3001" href="Relation.Nullary.Decidable.Core.html#2969" class="Bound">b?</a>
|
||
|
||
<a id="3005" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="3078" class="Comment">-- Relationship with booleans</a>
|
||
|
||
<a id="3109" class="Comment">-- `isYes` is a stricter version of `does`. The lack of computation</a>
|
||
<a id="3177" class="Comment">-- means that we can recover the proposition `P` from `isYes a?` by</a>
|
||
<a id="3245" class="Comment">-- unification. This is useful when we are using the decision procedure</a>
|
||
<a id="3317" class="Comment">-- for proof automation.</a>
|
||
|
||
<a id="isYes"></a><a id="3343" href="Relation.Nullary.Decidable.Core.html#3343" class="Function">isYes</a> <a id="3349" class="Symbol">:</a> <a id="3351" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="3355" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="3357" class="Symbol">→</a> <a id="3359" href="Agda.Builtin.Bool.html#173" class="Datatype">Bool</a>
|
||
<a id="3364" href="Relation.Nullary.Decidable.Core.html#3343" class="Function">isYes</a> <a id="3370" class="Symbol">(</a><a id="3371" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="3377" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="3385" class="Symbol">_)</a> <a id="3388" class="Symbol">=</a> <a id="3390" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a>
|
||
<a id="3395" href="Relation.Nullary.Decidable.Core.html#3343" class="Function">isYes</a> <a id="3401" class="Symbol">(</a><a id="3402" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="3408" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="3416" class="Symbol">_)</a> <a id="3419" class="Symbol">=</a> <a id="3421" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a>
|
||
|
||
<a id="isNo"></a><a id="3428" href="Relation.Nullary.Decidable.Core.html#3428" class="Function">isNo</a> <a id="3433" class="Symbol">:</a> <a id="3435" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="3439" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="3441" class="Symbol">→</a> <a id="3443" href="Agda.Builtin.Bool.html#173" class="Datatype">Bool</a>
|
||
<a id="3448" href="Relation.Nullary.Decidable.Core.html#3428" class="Function">isNo</a> <a id="3453" class="Symbol">=</a> <a id="3455" href="Data.Bool.Base.html#941" class="Function">not</a> <a id="3459" href="Function.Base.html#1115" class="Function Operator">∘</a> <a id="3461" href="Relation.Nullary.Decidable.Core.html#3343" class="Function">isYes</a>
|
||
|
||
<a id="True"></a><a id="3468" href="Relation.Nullary.Decidable.Core.html#3468" class="Function">True</a> <a id="3473" class="Symbol">:</a> <a id="3475" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="3479" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="3481" class="Symbol">→</a> <a id="3483" href="Agda.Primitive.html#388" class="Primitive">Set</a>
|
||
<a id="3487" href="Relation.Nullary.Decidable.Core.html#3468" class="Function">True</a> <a id="3492" class="Symbol">=</a> <a id="3494" href="Data.Bool.Base.html#1348" class="Function">T</a> <a id="3496" href="Function.Base.html#1115" class="Function Operator">∘</a> <a id="3498" href="Relation.Nullary.Decidable.Core.html#3343" class="Function">isYes</a>
|
||
|
||
<a id="False"></a><a id="3505" href="Relation.Nullary.Decidable.Core.html#3505" class="Function">False</a> <a id="3511" class="Symbol">:</a> <a id="3513" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="3517" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="3519" class="Symbol">→</a> <a id="3521" href="Agda.Primitive.html#388" class="Primitive">Set</a>
|
||
<a id="3525" href="Relation.Nullary.Decidable.Core.html#3505" class="Function">False</a> <a id="3531" class="Symbol">=</a> <a id="3533" href="Data.Bool.Base.html#1348" class="Function">T</a> <a id="3535" href="Function.Base.html#1115" class="Function Operator">∘</a> <a id="3537" href="Relation.Nullary.Decidable.Core.html#3428" class="Function">isNo</a>
|
||
|
||
<a id="3543" class="Comment">-- The traditional name for isYes is ⌊_⌋, indicating the stripping of evidence.</a>
|
||
<a id="⌊_⌋"></a><a id="3623" href="Relation.Nullary.Decidable.Core.html#3623" class="Function Operator">⌊_⌋</a> <a id="3627" class="Symbol">=</a> <a id="3629" href="Relation.Nullary.Decidable.Core.html#3343" class="Function">isYes</a>
|
||
|
||
<a id="3636" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="3709" class="Comment">-- Witnesses</a>
|
||
|
||
<a id="3723" class="Comment">-- Gives a witness to the "truth".</a>
|
||
<a id="toWitness"></a><a id="3758" href="Relation.Nullary.Decidable.Core.html#3758" class="Function">toWitness</a> <a id="3768" class="Symbol">:</a> <a id="3770" class="Symbol">{</a><a id="3771" href="Relation.Nullary.Decidable.Core.html#3771" class="Bound">a?</a> <a id="3774" class="Symbol">:</a> <a id="3776" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="3780" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="3781" class="Symbol">}</a> <a id="3783" class="Symbol">→</a> <a id="3785" href="Relation.Nullary.Decidable.Core.html#3468" class="Function">True</a> <a id="3790" href="Relation.Nullary.Decidable.Core.html#3771" class="Bound">a?</a> <a id="3793" class="Symbol">→</a> <a id="3795" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a>
|
||
<a id="3797" href="Relation.Nullary.Decidable.Core.html#3758" class="Function">toWitness</a> <a id="3807" class="Symbol">{</a><a id="3808" class="Argument">a?</a> <a id="3811" class="Symbol">=</a> <a id="3813" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="3819" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="3827" href="Relation.Nullary.Decidable.Core.html#3827" class="Bound">[a]</a><a id="3830" class="Symbol">}</a> <a id="3832" class="Symbol">_</a> <a id="3835" class="Symbol">=</a> <a id="3837" href="Relation.Nullary.Reflects.html#1582" class="Function">invert</a> <a id="3844" href="Relation.Nullary.Decidable.Core.html#3827" class="Bound">[a]</a>
|
||
<a id="3848" href="Relation.Nullary.Decidable.Core.html#3758" class="Function">toWitness</a> <a id="3858" class="Symbol">{</a><a id="3859" class="Argument">a?</a> <a id="3862" class="Symbol">=</a> <a id="3864" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="3870" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="3879" class="Symbol">_</a> <a id="3881" class="Symbol">}</a> <a id="3883" class="Symbol">()</a>
|
||
|
||
<a id="3887" class="Comment">-- Establishes a "truth", given a witness.</a>
|
||
<a id="fromWitness"></a><a id="3930" href="Relation.Nullary.Decidable.Core.html#3930" class="Function">fromWitness</a> <a id="3942" class="Symbol">:</a> <a id="3944" class="Symbol">{</a><a id="3945" href="Relation.Nullary.Decidable.Core.html#3945" class="Bound">a?</a> <a id="3948" class="Symbol">:</a> <a id="3950" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="3954" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="3955" class="Symbol">}</a> <a id="3957" class="Symbol">→</a> <a id="3959" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="3961" class="Symbol">→</a> <a id="3963" href="Relation.Nullary.Decidable.Core.html#3468" class="Function">True</a> <a id="3968" href="Relation.Nullary.Decidable.Core.html#3945" class="Bound">a?</a>
|
||
<a id="3971" href="Relation.Nullary.Decidable.Core.html#3930" class="Function">fromWitness</a> <a id="3983" class="Symbol">{</a><a id="3984" class="Argument">a?</a> <a id="3987" class="Symbol">=</a> <a id="3989" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="3995" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4005" class="Symbol">_</a> <a id="4007" class="Symbol">}</a> <a id="4009" class="Symbol">=</a> <a id="4011" href="Function.Base.html#725" class="Function">const</a> <a id="4017" class="Symbol">_</a>
|
||
<a id="4019" href="Relation.Nullary.Decidable.Core.html#3930" class="Function">fromWitness</a> <a id="4031" class="Symbol">{</a><a id="4032" class="Argument">a?</a> <a id="4035" class="Symbol">=</a> <a id="4037" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="4043" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4051" href="Relation.Nullary.Decidable.Core.html#4051" class="Bound">[¬a]</a><a id="4055" class="Symbol">}</a> <a id="4057" class="Symbol">=</a> <a id="4059" href="Relation.Nullary.Reflects.html#1582" class="Function">invert</a> <a id="4066" href="Relation.Nullary.Decidable.Core.html#4051" class="Bound">[¬a]</a>
|
||
|
||
<a id="4072" class="Comment">-- Variants for False.</a>
|
||
<a id="toWitnessFalse"></a><a id="4095" href="Relation.Nullary.Decidable.Core.html#4095" class="Function">toWitnessFalse</a> <a id="4110" class="Symbol">:</a> <a id="4112" class="Symbol">{</a><a id="4113" href="Relation.Nullary.Decidable.Core.html#4113" class="Bound">a?</a> <a id="4116" class="Symbol">:</a> <a id="4118" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="4122" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="4123" class="Symbol">}</a> <a id="4125" class="Symbol">→</a> <a id="4127" href="Relation.Nullary.Decidable.Core.html#3505" class="Function">False</a> <a id="4133" href="Relation.Nullary.Decidable.Core.html#4113" class="Bound">a?</a> <a id="4136" class="Symbol">→</a> <a id="4138" href="Relation.Nullary.Negation.Core.html#698" class="Function Operator">¬</a> <a id="4140" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a>
|
||
<a id="4142" href="Relation.Nullary.Decidable.Core.html#4095" class="Function">toWitnessFalse</a> <a id="4157" class="Symbol">{</a><a id="4158" class="Argument">a?</a> <a id="4161" class="Symbol">=</a> <a id="4163" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="4169" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4179" class="Symbol">_</a> <a id="4181" class="Symbol">}</a> <a id="4183" class="Symbol">()</a>
|
||
<a id="4186" href="Relation.Nullary.Decidable.Core.html#4095" class="Function">toWitnessFalse</a> <a id="4201" class="Symbol">{</a><a id="4202" class="Argument">a?</a> <a id="4205" class="Symbol">=</a> <a id="4207" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="4213" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4221" href="Relation.Nullary.Decidable.Core.html#4221" class="Bound">[¬a]</a><a id="4225" class="Symbol">}</a> <a id="4227" class="Symbol">_</a> <a id="4230" class="Symbol">=</a> <a id="4232" href="Relation.Nullary.Reflects.html#1582" class="Function">invert</a> <a id="4239" href="Relation.Nullary.Decidable.Core.html#4221" class="Bound">[¬a]</a>
|
||
|
||
<a id="fromWitnessFalse"></a><a id="4245" href="Relation.Nullary.Decidable.Core.html#4245" class="Function">fromWitnessFalse</a> <a id="4262" class="Symbol">:</a> <a id="4264" class="Symbol">{</a><a id="4265" href="Relation.Nullary.Decidable.Core.html#4265" class="Bound">a?</a> <a id="4268" class="Symbol">:</a> <a id="4270" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="4274" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="4275" class="Symbol">}</a> <a id="4277" class="Symbol">→</a> <a id="4279" href="Relation.Nullary.Negation.Core.html#698" class="Function Operator">¬</a> <a id="4281" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="4283" class="Symbol">→</a> <a id="4285" href="Relation.Nullary.Decidable.Core.html#3505" class="Function">False</a> <a id="4291" href="Relation.Nullary.Decidable.Core.html#4265" class="Bound">a?</a>
|
||
<a id="4294" href="Relation.Nullary.Decidable.Core.html#4245" class="Function">fromWitnessFalse</a> <a id="4311" class="Symbol">{</a><a id="4312" class="Argument">a?</a> <a id="4315" class="Symbol">=</a> <a id="4317" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="4323" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4331" href="Relation.Nullary.Decidable.Core.html#4331" class="Bound">[a]</a><a id="4334" class="Symbol">}</a> <a id="4336" class="Symbol">=</a> <a id="4338" href="Function.Base.html#1638" class="Function">flip</a> <a id="4343" href="Function.Base.html#1974" class="Function Operator">_$_</a> <a id="4347" class="Symbol">(</a><a id="4348" href="Relation.Nullary.Reflects.html#1582" class="Function">invert</a> <a id="4355" href="Relation.Nullary.Decidable.Core.html#4331" class="Bound">[a]</a><a id="4358" class="Symbol">)</a>
|
||
<a id="4360" href="Relation.Nullary.Decidable.Core.html#4245" class="Function">fromWitnessFalse</a> <a id="4377" class="Symbol">{</a><a id="4378" class="Argument">a?</a> <a id="4381" class="Symbol">=</a> <a id="4383" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="4389" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4398" class="Symbol">_</a> <a id="4400" class="Symbol">}</a> <a id="4402" class="Symbol">=</a> <a id="4404" href="Function.Base.html#725" class="Function">const</a> <a id="4410" class="Symbol">_</a>
|
||
|
||
<a id="4413" class="Comment">-- If a decision procedure returns "yes", then we can extract the</a>
|
||
<a id="4479" class="Comment">-- proof using from-yes.</a>
|
||
<a id="from-yes"></a><a id="4504" href="Relation.Nullary.Decidable.Core.html#4504" class="Function">from-yes</a> <a id="4513" class="Symbol">:</a> <a id="4515" class="Symbol">(</a><a id="4516" href="Relation.Nullary.Decidable.Core.html#4516" class="Bound">a?</a> <a id="4519" class="Symbol">:</a> <a id="4521" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="4525" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="4526" class="Symbol">)</a> <a id="4528" class="Symbol">→</a> <a id="4530" href="Relation.Nullary.Decidable.Core.html#1803" class="Function">From-yes</a> <a id="4539" href="Relation.Nullary.Decidable.Core.html#4516" class="Bound">a?</a>
|
||
<a id="4542" href="Relation.Nullary.Decidable.Core.html#4504" class="Function">from-yes</a> <a id="4551" class="Symbol">(</a><a id="4552" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="4558" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4566" href="Relation.Nullary.Decidable.Core.html#4566" class="Bound">[a]</a><a id="4569" class="Symbol">)</a> <a id="4571" class="Symbol">=</a> <a id="4573" href="Relation.Nullary.Reflects.html#1582" class="Function">invert</a> <a id="4580" href="Relation.Nullary.Decidable.Core.html#4566" class="Bound">[a]</a>
|
||
<a id="4584" href="Relation.Nullary.Decidable.Core.html#4504" class="Function">from-yes</a> <a id="4593" class="Symbol">(</a><a id="4594" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="4600" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4608" class="Symbol">_</a> <a id="4610" class="Symbol">)</a> <a id="4612" class="Symbol">=</a> <a id="4614" class="Symbol">_</a>
|
||
|
||
<a id="4617" class="Comment">-- If a decision procedure returns "no", then we can extract the proof</a>
|
||
<a id="4688" class="Comment">-- using from-no.</a>
|
||
<a id="from-no"></a><a id="4706" href="Relation.Nullary.Decidable.Core.html#4706" class="Function">from-no</a> <a id="4714" class="Symbol">:</a> <a id="4716" class="Symbol">(</a><a id="4717" href="Relation.Nullary.Decidable.Core.html#4717" class="Bound">a?</a> <a id="4720" class="Symbol">:</a> <a id="4722" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="4726" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="4727" class="Symbol">)</a> <a id="4729" class="Symbol">→</a> <a id="4731" href="Relation.Nullary.Decidable.Core.html#1904" class="Function">From-no</a> <a id="4739" href="Relation.Nullary.Decidable.Core.html#4717" class="Bound">a?</a>
|
||
<a id="4742" href="Relation.Nullary.Decidable.Core.html#4706" class="Function">from-no</a> <a id="4750" class="Symbol">(</a><a id="4751" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="4757" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4765" href="Relation.Nullary.Decidable.Core.html#4765" class="Bound">[¬a]</a><a id="4769" class="Symbol">)</a> <a id="4771" class="Symbol">=</a> <a id="4773" href="Relation.Nullary.Reflects.html#1582" class="Function">invert</a> <a id="4780" href="Relation.Nullary.Decidable.Core.html#4765" class="Bound">[¬a]</a>
|
||
<a id="4785" href="Relation.Nullary.Decidable.Core.html#4706" class="Function">from-no</a> <a id="4793" class="Symbol">(</a><a id="4794" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="4800" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="4810" class="Symbol">_</a> <a id="4812" class="Symbol">)</a> <a id="4814" class="Symbol">=</a> <a id="4816" class="Symbol">_</a>
|
||
|
||
<a id="4819" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="4892" class="Comment">-- Maps</a>
|
||
|
||
<a id="map′"></a><a id="4901" href="Relation.Nullary.Decidable.Core.html#4901" class="Function">map′</a> <a id="4906" class="Symbol">:</a> <a id="4908" class="Symbol">(</a><a id="4909" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="4911" class="Symbol">→</a> <a id="4913" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a><a id="4914" class="Symbol">)</a> <a id="4916" class="Symbol">→</a> <a id="4918" class="Symbol">(</a><a id="4919" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a> <a id="4921" class="Symbol">→</a> <a id="4923" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="4924" class="Symbol">)</a> <a id="4926" class="Symbol">→</a> <a id="4928" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="4932" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="4934" class="Symbol">→</a> <a id="4936" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="4940" href="Relation.Nullary.Decidable.Core.html#915" class="Generalizable">B</a>
|
||
<a id="4942" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="4948" class="Symbol">(</a><a id="4949" href="Relation.Nullary.Decidable.Core.html#4901" class="Function">map′</a> <a id="4954" href="Relation.Nullary.Decidable.Core.html#4954" class="Bound">A→B</a> <a id="4958" href="Relation.Nullary.Decidable.Core.html#4958" class="Bound">B→A</a> <a id="4962" href="Relation.Nullary.Decidable.Core.html#4962" class="Bound">a?</a><a id="4964" class="Symbol">)</a> <a id="4984" class="Symbol">=</a> <a id="4986" href="Relation.Nullary.Decidable.Core.html#1551" class="Field">does</a> <a id="4991" href="Relation.Nullary.Decidable.Core.html#4962" class="Bound">a?</a>
|
||
<a id="4994" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="5000" class="Symbol">(</a><a id="5001" href="Relation.Nullary.Decidable.Core.html#4901" class="Function">map′</a> <a id="5006" href="Relation.Nullary.Decidable.Core.html#5006" class="Bound">A→B</a> <a id="5010" href="Relation.Nullary.Decidable.Core.html#5010" class="Bound">B→A</a> <a id="5014" class="Symbol">(</a><a id="5015" href="Agda.Builtin.Bool.html#198" class="InductiveConstructor">true</a> <a id="5021" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="5030" href="Relation.Nullary.Decidable.Core.html#5030" class="Bound">[a]</a><a id="5033" class="Symbol">))</a> <a id="5036" class="Symbol">=</a> <a id="5038" href="Relation.Nullary.Reflects.html#982" class="InductiveConstructor">ofʸ</a> <a id="5042" class="Symbol">(</a><a id="5043" href="Relation.Nullary.Decidable.Core.html#5006" class="Bound">A→B</a> <a id="5047" class="Symbol">(</a><a id="5048" href="Relation.Nullary.Reflects.html#1582" class="Function">invert</a> <a id="5055" href="Relation.Nullary.Decidable.Core.html#5030" class="Bound">[a]</a><a id="5058" class="Symbol">))</a>
|
||
<a id="5061" href="Relation.Nullary.Decidable.Core.html#1568" class="Field">proof</a> <a id="5067" class="Symbol">(</a><a id="5068" href="Relation.Nullary.Decidable.Core.html#4901" class="Function">map′</a> <a id="5073" href="Relation.Nullary.Decidable.Core.html#5073" class="Bound">A→B</a> <a id="5077" href="Relation.Nullary.Decidable.Core.html#5077" class="Bound">B→A</a> <a id="5081" class="Symbol">(</a><a id="5082" href="Agda.Builtin.Bool.html#192" class="InductiveConstructor">false</a> <a id="5088" href="Relation.Nullary.Decidable.Core.html#1529" class="InductiveConstructor Operator">because</a> <a id="5096" href="Relation.Nullary.Decidable.Core.html#5096" class="Bound">[¬a]</a><a id="5100" class="Symbol">))</a> <a id="5103" class="Symbol">=</a> <a id="5105" href="Relation.Nullary.Reflects.html#1019" class="InductiveConstructor">ofⁿ</a> <a id="5109" class="Symbol">(</a><a id="5110" href="Relation.Nullary.Reflects.html#1582" class="Function">invert</a> <a id="5117" href="Relation.Nullary.Decidable.Core.html#5096" class="Bound">[¬a]</a> <a id="5122" href="Function.Base.html#1115" class="Function Operator">∘</a> <a id="5124" href="Relation.Nullary.Decidable.Core.html#5077" class="Bound">B→A</a><a id="5127" class="Symbol">)</a>
|
||
|
||
<a id="5130" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="5203" class="Comment">-- Relationship with double-negation</a>
|
||
|
||
<a id="5241" class="Comment">-- Decidable predicates are stable.</a>
|
||
|
||
<a id="decidable-stable"></a><a id="5278" href="Relation.Nullary.Decidable.Core.html#5278" class="Function">decidable-stable</a> <a id="5295" class="Symbol">:</a> <a id="5297" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="5301" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a> <a id="5303" class="Symbol">→</a> <a id="5305" href="Relation.Nullary.Negation.Core.html#930" class="Function">Stable</a> <a id="5312" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a>
|
||
<a id="5314" href="Relation.Nullary.Decidable.Core.html#5278" class="Function">decidable-stable</a> <a id="5331" class="Symbol">(</a><a id="5332" href="Relation.Nullary.Decidable.Core.html#1618" class="InductiveConstructor">yes</a> <a id="5336" href="Relation.Nullary.Decidable.Core.html#5336" class="Bound">a</a><a id="5337" class="Symbol">)</a> <a id="5339" href="Relation.Nullary.Decidable.Core.html#5339" class="Bound">¬¬a</a> <a id="5343" class="Symbol">=</a> <a id="5345" href="Relation.Nullary.Decidable.Core.html#5336" class="Bound">a</a>
|
||
<a id="5347" href="Relation.Nullary.Decidable.Core.html#5278" class="Function">decidable-stable</a> <a id="5364" class="Symbol">(</a><a id="5365" href="Relation.Nullary.Decidable.Core.html#1655" class="InductiveConstructor">no</a> <a id="5368" href="Relation.Nullary.Decidable.Core.html#5368" class="Bound">¬a</a><a id="5370" class="Symbol">)</a> <a id="5372" href="Relation.Nullary.Decidable.Core.html#5372" class="Bound">¬¬a</a> <a id="5376" class="Symbol">=</a> <a id="5378" href="Data.Empty.Irrelevant.html#336" class="Function">⊥-elim</a> <a id="5385" class="Symbol">(</a><a id="5386" href="Relation.Nullary.Decidable.Core.html#5372" class="Bound">¬¬a</a> <a id="5390" href="Relation.Nullary.Decidable.Core.html#5368" class="Bound">¬a</a><a id="5392" class="Symbol">)</a>
|
||
|
||
<a id="¬-drop-Dec"></a><a id="5395" href="Relation.Nullary.Decidable.Core.html#5395" class="Function">¬-drop-Dec</a> <a id="5406" class="Symbol">:</a> <a id="5408" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="5412" class="Symbol">(</a><a id="5413" href="Relation.Nullary.Negation.Core.html#698" class="Function Operator">¬</a> <a id="5415" href="Relation.Nullary.Negation.Core.html#698" class="Function Operator">¬</a> <a id="5417" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="5418" class="Symbol">)</a> <a id="5420" class="Symbol">→</a> <a id="5422" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="5426" class="Symbol">(</a><a id="5427" href="Relation.Nullary.Negation.Core.html#698" class="Function Operator">¬</a> <a id="5429" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="5430" class="Symbol">)</a>
|
||
<a id="5432" href="Relation.Nullary.Decidable.Core.html#5395" class="Function">¬-drop-Dec</a> <a id="5443" href="Relation.Nullary.Decidable.Core.html#5443" class="Bound">¬¬a?</a> <a id="5448" class="Symbol">=</a> <a id="5450" href="Relation.Nullary.Decidable.Core.html#4901" class="Function">map′</a> <a id="5455" href="Relation.Nullary.Negation.Core.html#1711" class="Function">negated-stable</a> <a id="5470" href="Relation.Nullary.Negation.Core.html#1226" class="Function">contradiction</a> <a id="5484" class="Symbol">(</a><a id="5485" href="Relation.Nullary.Decidable.Core.html#2517" class="Function">¬?</a> <a id="5488" href="Relation.Nullary.Decidable.Core.html#5443" class="Bound">¬¬a?</a><a id="5492" class="Symbol">)</a>
|
||
|
||
<a id="5495" class="Comment">-- A double-negation-translated variant of excluded middle (or: every</a>
|
||
<a id="5565" class="Comment">-- nullary relation is decidable in the double-negation monad).</a>
|
||
|
||
<a id="¬¬-excluded-middle"></a><a id="5630" href="Relation.Nullary.Decidable.Core.html#5630" class="Function">¬¬-excluded-middle</a> <a id="5649" class="Symbol">:</a> <a id="5651" href="Relation.Nullary.Negation.Core.html#837" class="Function">DoubleNegation</a> <a id="5666" class="Symbol">(</a><a id="5667" href="Relation.Nullary.Decidable.Core.html#1485" class="Record">Dec</a> <a id="5671" href="Relation.Nullary.Decidable.Core.html#913" class="Generalizable">A</a><a id="5672" class="Symbol">)</a>
|
||
<a id="5674" href="Relation.Nullary.Decidable.Core.html#5630" class="Function">¬¬-excluded-middle</a> <a id="5693" href="Relation.Nullary.Decidable.Core.html#5693" class="Bound">¬?a</a> <a id="5697" class="Symbol">=</a> <a id="5699" href="Relation.Nullary.Decidable.Core.html#5693" class="Bound">¬?a</a> <a id="5703" class="Symbol">(</a><a id="5704" href="Relation.Nullary.Decidable.Core.html#1655" class="InductiveConstructor">no</a> <a id="5707" class="Symbol">(λ</a> <a id="5710" href="Relation.Nullary.Decidable.Core.html#5710" class="Bound">a</a> <a id="5712" class="Symbol">→</a> <a id="5714" href="Relation.Nullary.Decidable.Core.html#5693" class="Bound">¬?a</a> <a id="5718" class="Symbol">(</a><a id="5719" href="Relation.Nullary.Decidable.Core.html#1618" class="InductiveConstructor">yes</a> <a id="5723" href="Relation.Nullary.Decidable.Core.html#5710" class="Bound">a</a><a id="5724" class="Symbol">)))</a>
|
||
|
||
|
||
<a id="5730" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="5803" class="Comment">-- DEPRECATED NAMES</a>
|
||
<a id="5823" class="Comment">------------------------------------------------------------------------</a>
|
||
<a id="5896" class="Comment">-- Please use the new names as continuing support for the old names is</a>
|
||
<a id="5967" class="Comment">-- not guaranteed.</a>
|
||
|
||
<a id="5987" class="Comment">-- Version 2.0</a>
|
||
|
||
<a id="excluded-middle"></a><a id="6003" href="Relation.Nullary.Decidable.Core.html#6003" class="Function">excluded-middle</a> <a id="6019" class="Symbol">=</a> <a id="6021" href="Relation.Nullary.Decidable.Core.html#5630" class="Function">¬¬-excluded-middle</a>
|
||
<a id="6040" class="Symbol">{-#</a> <a id="6044" class="Keyword">WARNING_ON_USAGE</a> <a id="6061" class="Pragma">excluded-middle</a>
|
||
<a id="6077" class="String">"Warning: excluded-middle was deprecated in v2.0.
|
||
Please use ¬¬-excluded-middle instead."</a>
|
||
<a id="6167" class="Symbol">#-}</a>
|
||
</pre></body></html> |