mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
144 lines
64 KiB
HTML
144 lines
64 KiB
HTML
|
<!DOCTYPE html>
|
|||
|
<html xmlns="http://www.w3.org/1999/xhtml" lang="" xml:lang="">
|
|||
|
<head>
|
|||
|
<meta charset="utf-8" />
|
|||
|
<meta name="generator" content="pandoc" />
|
|||
|
<meta name="viewport" content="width=device-width, initial-scale=1.0, user-scalable=yes" />
|
|||
|
<title>Category.Ambient.Setoids</title>
|
|||
|
<style>
|
|||
|
code{white-space: pre-wrap;}
|
|||
|
span.smallcaps{font-variant: small-caps;}
|
|||
|
div.columns{display: flex; gap: min(4vw, 1.5em);}
|
|||
|
div.column{flex: auto; overflow-x: auto;}
|
|||
|
div.hanging-indent{margin-left: 1.5em; text-indent: -1.5em;}
|
|||
|
/* The extra [class] is a hack that increases specificity enough to
|
|||
|
override a similar rule in reveal.js */
|
|||
|
ul.task-list[class]{list-style: none;}
|
|||
|
ul.task-list li input[type="checkbox"] {
|
|||
|
font-size: inherit;
|
|||
|
width: 0.8em;
|
|||
|
margin: 0 0.8em 0.2em -1.6em;
|
|||
|
vertical-align: middle;
|
|||
|
}
|
|||
|
</style>
|
|||
|
<link rel="stylesheet" href="Agda.css" />
|
|||
|
<!--[if lt IE 9]>
|
|||
|
<script src="//cdnjs.cloudflare.com/ajax/libs/html5shiv/3.7.3/html5shiv-printshiv.min.js"></script>
|
|||
|
<![endif]-->
|
|||
|
</head>
|
|||
|
<body>
|
|||
|
<!--
|
|||
|
<pre class="Agda"><a id="14" class="Keyword">open</a> <a id="19" class="Keyword">import</a> <a id="26" href="Level.html" class="Module">Level</a> <a id="32" class="Keyword">using</a> <a id="38" class="Symbol">(</a><a id="39" href="Level.html#466" class="InductiveConstructor">lift</a><a id="43" class="Symbol">;</a> <a id="45" href="Level.html#409" class="Record">Lift</a><a id="49" class="Symbol">)</a> <a id="51" class="Keyword">renaming</a> <a id="60" class="Symbol">(</a><a id="61" href="Agda.Primitive.html#931" class="Primitive">suc</a> <a id="65" class="Symbol">to</a> <a id="68" class="Primitive">ℓ-suc</a><a id="73" class="Symbol">)</a>
|
|||
|
<a id="75" class="Keyword">open</a> <a id="80" class="Keyword">import</a> <a id="87" href="Category.Ambient.html" class="Module">Category.Ambient</a>
|
|||
|
<a id="104" class="Keyword">open</a> <a id="109" class="Keyword">import</a> <a id="116" href="Categories.Category.Instance.Setoids.html" class="Module">Categories.Category.Instance.Setoids</a>
|
|||
|
<a id="153" class="Keyword">open</a> <a id="158" class="Keyword">import</a> <a id="165" href="Categories.Category.Instance.Properties.Setoids.Extensive.html" class="Module">Categories.Category.Instance.Properties.Setoids.Extensive</a>
|
|||
|
<a id="223" class="Keyword">open</a> <a id="228" class="Keyword">import</a> <a id="235" href="Categories.Category.Instance.Properties.Setoids.CCC.html" class="Module">Categories.Category.Instance.Properties.Setoids.CCC</a>
|
|||
|
<a id="287" class="Keyword">open</a> <a id="292" class="Keyword">import</a> <a id="299" href="Categories.Category.Monoidal.Instance.Setoids.html" class="Module">Categories.Category.Monoidal.Instance.Setoids</a>
|
|||
|
<a id="345" class="Keyword">open</a> <a id="350" class="Keyword">import</a> <a id="357" href="Categories.Category.Instance.SingletonSet.html" class="Module">Categories.Category.Instance.SingletonSet</a>
|
|||
|
<a id="399" class="Keyword">open</a> <a id="404" class="Keyword">import</a> <a id="411" href="Categories.Object.NaturalNumbers.html" class="Module">Categories.Object.NaturalNumbers</a>
|
|||
|
<a id="444" class="Keyword">open</a> <a id="449" class="Keyword">import</a> <a id="456" href="Relation.Binary.html" class="Module">Relation.Binary</a>
|
|||
|
<a id="472" class="Keyword">open</a> <a id="477" class="Keyword">import</a> <a id="484" href="Data.Unit.Polymorphic.html" class="Module">Data.Unit.Polymorphic</a> <a id="506" class="Keyword">using</a> <a id="512" class="Symbol">(</a><a id="513" href="Data.Unit.Polymorphic.Base.html#533" class="Function">tt</a><a id="515" class="Symbol">;</a> <a id="517" href="Data.Unit.Polymorphic.Base.html#489" class="Function">⊤</a><a id="518" class="Symbol">)</a>
|
|||
|
<a id="520" class="Keyword">open</a> <a id="525" class="Keyword">import</a> <a id="532" href="Data.Empty.Polymorphic.html" class="Module">Data.Empty.Polymorphic</a> <a id="555" class="Keyword">using</a> <a id="561" class="Symbol">(</a><a id="562" href="Data.Empty.Polymorphic.html#340" class="Function">⊥</a><a id="563" class="Symbol">)</a>
|
|||
|
<a id="565" class="Keyword">open</a> <a id="570" class="Keyword">import</a> <a id="577" href="Function.Bundles.html" class="Module">Function.Bundles</a> <a id="594" class="Keyword">using</a> <a id="600" class="Symbol">(</a><a id="601" href="Function.Bundles.html#15133" class="Function Operator">_⟨$⟩_</a><a id="606" class="Symbol">)</a> <a id="608" class="Keyword">renaming</a> <a id="617" class="Symbol">(</a><a id="618" href="Function.Bundles.html#2043" class="Record">Func</a> <a id="623" class="Symbol">to</a> <a id="626" class="Record">_⟶_</a><a id="629" class="Symbol">)</a>
|
|||
|
<a id="631" class="Keyword">open</a> <a id="636" class="Keyword">import</a> <a id="643" href="Function.Construct.Identity.html" class="Module">Function.Construct.Identity</a> <a id="671" class="Symbol">as</a> <a id="674" class="Module">idₛ</a>
|
|||
|
<a id="678" class="Keyword">open</a> <a id="683" class="Keyword">import</a> <a id="690" href="Function.Construct.Setoid.html" class="Module">Function.Construct.Setoid</a> <a id="716" class="Keyword">renaming</a> <a id="725" class="Symbol">(</a><a id="726" href="Function.Construct.Setoid.html#376" class="Function">setoid</a> <a id="733" class="Symbol">to</a> <a id="736" class="Function">_⇨_</a><a id="739" class="Symbol">;</a> <a id="741" href="Function.Construct.Setoid.html#888" class="Function Operator">_∙_</a> <a id="745" class="Symbol">to</a> <a id="748" class="Function Operator">_∘_</a><a id="751" class="Symbol">)</a>
|
|||
|
<a id="753" class="Keyword">open</a> <a id="758" class="Keyword">import</a> <a id="765" href="Categories.Object.NaturalNumbers.Properties.Parametrized.html" class="Module">Categories.Object.NaturalNumbers.Properties.Parametrized</a>
|
|||
|
<a id="822" class="Keyword">open</a> <a id="827" class="Keyword">import</a> <a id="834" href="Categories.Category.Cocartesian.html" class="Module">Categories.Category.Cocartesian</a>
|
|||
|
<a id="866" class="Keyword">import</a> <a id="873" href="Relation.Binary.Reasoning.Setoid.html" class="Module">Relation.Binary.Reasoning.Setoid</a> <a id="906" class="Symbol">as</a> <a id="909" class="Module">SetoidR</a>
|
|||
|
|
|||
|
<a id="918" class="Keyword">import</a> <a id="925" href="Relation.Binary.PropositionalEquality.html" class="Module">Relation.Binary.PropositionalEquality</a> <a id="963" class="Symbol">as</a> <a id="966" class="Module">Eq</a>
|
|||
|
<a id="969" class="Keyword">open</a> <a id="974" href="Relation.Binary.PropositionalEquality.html" class="Module">Eq</a> <a id="977" class="Keyword">using</a> <a id="983" class="Symbol">(</a><a id="984" href="Agda.Builtin.Equality.html#150" class="Datatype Operator">_≡_</a><a id="987" class="Symbol">)</a>
|
|||
|
</pre>-->
|
|||
|
<h1
|
|||
|
id="the-category-of-setoids-can-be-used-as-instance-for-ambient-category">The
|
|||
|
category of Setoids can be used as instance for ambient category</h1>
|
|||
|
<p>Most of the required properties are already proven in the
|
|||
|
agda-categories library, we are only left to construct the natural
|
|||
|
numbers object.</p>
|
|||
|
<pre class="Agda"><a id="1220" class="Keyword">module</a> <a id="1227" href="Category.Ambient.Setoids.html" class="Module">Category.Ambient.Setoids</a> <a id="1252" class="Symbol">{</a><a id="1253" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="1254" class="Symbol">}</a> <a id="1256" class="Keyword">where</a>
|
|||
|
<a id="1264" class="Keyword">open</a> <a id="1269" href="Category.Ambient.Setoids.html#626" class="Module">_⟶_</a> <a id="1273" class="Keyword">using</a> <a id="1279" class="Symbol">(</a><a id="1280" href="Function.Bundles.html#2113" class="Field">cong</a><a id="1284" class="Symbol">)</a>
|
|||
|
|
|||
|
<a id="1289" class="Comment">-- equality on setoid functions</a>
|
|||
|
<a id="1323" class="Keyword">private</a>
|
|||
|
<a id="_≋_"></a><a id="1335" href="Category.Ambient.Setoids.html#1335" class="Function Operator">_≋_</a> <a id="1339" class="Symbol">:</a> <a id="1341" class="Symbol">∀</a> <a id="1343" class="Symbol">{</a><a id="1344" href="Category.Ambient.Setoids.html#1344" class="Bound">A</a> <a id="1346" href="Category.Ambient.Setoids.html#1346" class="Bound">B</a> <a id="1348" class="Symbol">:</a> <a id="1350" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="1357" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="1359" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="1360" class="Symbol">}</a> <a id="1362" class="Symbol">→</a> <a id="1364" href="Category.Ambient.Setoids.html#1344" class="Bound">A</a> <a id="1366" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="1368" href="Category.Ambient.Setoids.html#1346" class="Bound">B</a> <a id="1370" class="Symbol">→</a> <a id="1372" href="Category.Ambient.Setoids.html#1344" class="Bound">A</a> <a id="1374" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="1376" href="Category.Ambient.Setoids.html#1346" class="Bound">B</a> <a id="1378" class="Symbol">→</a> <a id="1380" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1384" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a>
|
|||
|
<a id="1390" href="Category.Ambient.Setoids.html#1335" class="Function Operator">_≋_</a> <a id="1394" class="Symbol">{</a><a id="1395" href="Category.Ambient.Setoids.html#1395" class="Bound">A</a><a id="1396" class="Symbol">}</a> <a id="1398" class="Symbol">{</a><a id="1399" href="Category.Ambient.Setoids.html#1399" class="Bound">B</a><a id="1400" class="Symbol">}</a> <a id="1402" href="Category.Ambient.Setoids.html#1402" class="Bound">f</a> <a id="1404" href="Category.Ambient.Setoids.html#1404" class="Bound">g</a> <a id="1406" class="Symbol">=</a> <a id="1408" href="Relation.Binary.Bundles.html#1169" class="Field Operator">Setoid._≈_</a> <a id="1419" class="Symbol">(</a><a id="1420" href="Category.Ambient.Setoids.html#1395" class="Bound">A</a> <a id="1422" href="Category.Ambient.Setoids.html#736" class="Function Operator">⇨</a> <a id="1424" href="Category.Ambient.Setoids.html#1399" class="Bound">B</a><a id="1425" class="Symbol">)</a> <a id="1427" href="Category.Ambient.Setoids.html#1402" class="Bound">f</a> <a id="1429" href="Category.Ambient.Setoids.html#1404" class="Bound">g</a>
|
|||
|
<a id="≋-sym"></a><a id="1435" href="Category.Ambient.Setoids.html#1435" class="Function">≋-sym</a> <a id="1441" class="Symbol">:</a> <a id="1443" class="Symbol">∀</a> <a id="1445" class="Symbol">{</a><a id="1446" href="Category.Ambient.Setoids.html#1446" class="Bound">A</a> <a id="1448" href="Category.Ambient.Setoids.html#1448" class="Bound">B</a> <a id="1450" class="Symbol">:</a> <a id="1452" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="1459" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="1461" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="1462" class="Symbol">}</a> <a id="1464" class="Symbol">{</a><a id="1465" href="Category.Ambient.Setoids.html#1465" class="Bound">f</a> <a id="1467" href="Category.Ambient.Setoids.html#1467" class="Bound">g</a> <a id="1469" class="Symbol">:</a> <a id="1471" href="Category.Ambient.Setoids.html#1446" class="Bound">A</a> <a id="1473" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="1475" href="Category.Ambient.Setoids.html#1448" class="Bound">B</a><a id="1476" class="Symbol">}</a> <a id="1478" class="Symbol">→</a> <a id="1480" href="Category.Ambient.Setoids.html#1465" class="Bound">f</a> <a id="1482" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="1484" href="Category.Ambient.Setoids.html#1467" class="Bound">g</a> <a id="1486" class="Symbol">→</a> <a id="1488" href="Category.Ambient.Setoids.html#1467" class="Bound">g</a> <a id="1490" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="1492" href="Category.Ambient.Setoids.html#1465" class="Bound">f</a>
|
|||
|
<a id="1498" href="Category.Ambient.Setoids.html#1435" class="Function">≋-sym</a> <a id="1504" class="Symbol">{</a><a id="1505" href="Category.Ambient.Setoids.html#1505" class="Bound">A</a><a id="1506" class="Symbol">}</a> <a id="1508" class="Symbol">{</a><a id="1509" href="Category.Ambient.Setoids.html#1509" class="Bound">B</a><a id="1510" class="Symbol">}</a> <a id="1512" class="Symbol">{</a><a id="1513" href="Category.Ambient.Setoids.html#1513" class="Bound">f</a><a id="1514" class="Symbol">}</a> <a id="1516" class="Symbol">{</a><a id="1517" href="Category.Ambient.Setoids.html#1517" class="Bound">g</a><a id="1518" class="Symbol">}</a> <a id="1520" class="Symbol">=</a> <a id="1522" href="Relation.Binary.Structures.html#1622" class="Field">IsEquivalence.sym</a> <a id="1540" class="Symbol">(</a><a id="1541" href="Relation.Binary.Bundles.html#1203" class="Field">Setoid.isEquivalence</a> <a id="1562" class="Symbol">(</a><a id="1563" href="Category.Ambient.Setoids.html#1505" class="Bound">A</a> <a id="1565" href="Category.Ambient.Setoids.html#736" class="Function Operator">⇨</a> <a id="1567" href="Category.Ambient.Setoids.html#1509" class="Bound">B</a><a id="1568" class="Symbol">))</a> <a id="1571" class="Symbol">{</a><a id="1572" href="Category.Ambient.Setoids.html#1513" class="Bound">f</a><a id="1573" class="Symbol">}</a> <a id="1575" class="Symbol">{</a><a id="1576" href="Category.Ambient.Setoids.html#1517" class="Bound">g</a><a id="1577" class="Symbol">}</a>
|
|||
|
<a id="≋-trans"></a><a id="1583" href="Category.Ambient.Setoids.html#1583" class="Function">≋-trans</a> <a id="1591" class="Symbol">:</a> <a id="1593" class="Symbol">∀</a> <a id="1595" class="Symbol">{</a><a id="1596" href="Category.Ambient.Setoids.html#1596" class="Bound">A</a> <a id="1598" href="Category.Ambient.Setoids.html#1598" class="Bound">B</a> <a id="1600" class="Symbol">:</a> <a id="1602" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="1609" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="1611" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="1612" class="Symbol">}</a> <a id="1614" class="Symbol">{</a><a id="1615" href="Category.Ambient.Setoids.html#1615" class="Bound">f</a> <a id="1617" href="Category.Ambient.Setoids.html#1617" class="Bound">g</a> <a id="1619" href="Category.Ambient.Setoids.html#1619" class="Bound">h</a> <a id="1621" class="Symbol">:</a> <a id="1623" href="Category.Ambient.Setoids.html#1596" class="Bound">A</a> <a id="1625" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="1627" href="Category.Ambient.Setoids.html#1598" class="Bound">B</a><a id="1628" class="Symbol">}</a> <a id="1630" class="Symbol">→</a> <a id="1632" href="Category.Ambient.Setoids.html#1615" class="Bound">f</a> <a id="1634" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="1636" href="Category.Ambient.Setoids.html#1617" class="Bound">g</a> <a id="1638" class="Symbol">→</a> <a id="1640" href="Category.Ambient.Setoids.html#1617" class="Bound">g</a> <a id="1642" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="1644" href="Category.Ambient.Setoids.html#1619" class="Bound">h</a> <a id="1646" class="Symbol">→</a> <a id="1648" href="Category.Ambient.Setoids.html#1615" class="Bound">f</a> <a id="1650" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="1652" href="Category.Ambient.Setoids.html#1619" class="Bound">h</a>
|
|||
|
<a id="1658" href="Category.Ambient.Setoids.html#1583" class="Function">≋-trans</a> <a id="1666" class="Symbol">{</a><a id="1667" href="Category.Ambient.Setoids.html#1667" class="Bound">A</a><a id="1668" class="Symbol">}</a> <a id="1670" class="Symbol">{</a><a id="1671" href="Category.Ambient.Setoids.html#1671" class="Bound">B</a><a id="1672" class="Symbol">}</a> <a id="1674" class="Symbol">{</a><a id="1675" href="Category.Ambient.Setoids.html#1675" class="Bound">f</a><a id="1676" class="Symbol">}</a> <a id="1678" class="Symbol">{</a><a id="1679" href="Category.Ambient.Setoids.html#1679" class="Bound">g</a><a id="1680" class="Symbol">}</a> <a id="1682" class="Symbol">{</a><a id="1683" href="Category.Ambient.Setoids.html#1683" class="Bound">h</a><a id="1684" class="Symbol">}</a> <a id="1686" class="Symbol">=</a> <a id="1688" href="Relation.Binary.Structures.html#1648" class="Field">IsEquivalence.trans</a> <a id="1708" class="Symbol">(</a><a id="1709" href="Relation.Binary.Bundles.html#1203" class="Field">Setoid.isEquivalence</a> <a id="1730" class="Symbol">(</a><a id="1731" href="Category.Ambient.Setoids.html#1667" class="Bound">A</a> <a id="1733" href="Category.Ambient.Setoids.html#736" class="Function Operator">⇨</a> <a id="1735" href="Category.Ambient.Setoids.html#1671" class="Bound">B</a><a id="1736" class="Symbol">))</a> <a id="1739" class="Symbol">{</a><a id="1740" href="Category.Ambient.Setoids.html#1675" class="Bound">f</a><a id="1741" class="Symbol">}</a> <a id="1743" class="Symbol">{</a><a id="1744" href="Category.Ambient.Setoids.html#1679" class="Bound">g</a><a id="1745" class="Symbol">}</a> <a id="1747" class="Symbol">{</a><a id="1748" href="Category.Ambient.Setoids.html#1683" class="Bound">h</a><a id="1749" class="Symbol">}</a>
|
|||
|
|
|||
|
<a id="1754" class="Comment">-- we define ℕ ourselves, instead of importing it, to avoid lifting the universe levels (builtin Nats are defined on Set₀)</a>
|
|||
|
<a id="1879" class="Keyword">data</a> <a id="ℕ"></a><a id="1884" href="Category.Ambient.Setoids.html#1884" class="Datatype">ℕ</a> <a id="1886" class="Symbol">:</a> <a id="1888" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1892" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="1894" class="Keyword">where</a>
|
|||
|
<a id="ℕ.zero"></a><a id="1904" href="Category.Ambient.Setoids.html#1904" class="InductiveConstructor">zero</a> <a id="1909" class="Symbol">:</a> <a id="1911" href="Category.Ambient.Setoids.html#1884" class="Datatype">ℕ</a>
|
|||
|
<a id="ℕ.suc"></a><a id="1917" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="1921" class="Symbol">:</a> <a id="1923" href="Category.Ambient.Setoids.html#1884" class="Datatype">ℕ</a> <a id="1925" class="Symbol">→</a> <a id="1927" href="Category.Ambient.Setoids.html#1884" class="Datatype">ℕ</a>
|
|||
|
|
|||
|
<a id="suc-cong"></a><a id="1932" href="Category.Ambient.Setoids.html#1932" class="Function">suc-cong</a> <a id="1941" class="Symbol">:</a> <a id="1943" class="Symbol">∀</a> <a id="1945" class="Symbol">{</a><a id="1946" href="Category.Ambient.Setoids.html#1946" class="Bound">n</a> <a id="1948" href="Category.Ambient.Setoids.html#1948" class="Bound">m</a><a id="1949" class="Symbol">}</a> <a id="1951" class="Symbol">→</a> <a id="1953" href="Category.Ambient.Setoids.html#1946" class="Bound">n</a> <a id="1955" href="Agda.Builtin.Equality.html#150" class="Datatype Operator">≡</a> <a id="1957" href="Category.Ambient.Setoids.html#1948" class="Bound">m</a> <a id="1959" class="Symbol">→</a> <a id="1961" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="1965" href="Category.Ambient.Setoids.html#1946" class="Bound">n</a> <a id="1967" href="Agda.Builtin.Equality.html#150" class="Datatype Operator">≡</a> <a id="1969" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="1973" href="Category.Ambient.Setoids.html#1948" class="Bound">m</a>
|
|||
|
<a id="1977" href="Category.Ambient.Setoids.html#1932" class="Function">suc-cong</a> <a id="1986" href="Category.Ambient.Setoids.html#1986" class="Bound">n≡m</a> <a id="1990" class="Keyword">rewrite</a> <a id="1998" href="Category.Ambient.Setoids.html#1986" class="Bound">n≡m</a> <a id="2002" class="Symbol">=</a> <a id="2004" href="Agda.Builtin.Equality.html#207" class="InductiveConstructor">Eq.refl</a>
|
|||
|
|
|||
|
<a id="suc-inj"></a><a id="2015" href="Category.Ambient.Setoids.html#2015" class="Function">suc-inj</a> <a id="2023" class="Symbol">:</a> <a id="2025" class="Symbol">∀</a> <a id="2027" class="Symbol">{</a><a id="2028" href="Category.Ambient.Setoids.html#2028" class="Bound">n</a> <a id="2030" href="Category.Ambient.Setoids.html#2030" class="Bound">m</a><a id="2031" class="Symbol">}</a> <a id="2033" class="Symbol">→</a> <a id="2035" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="2039" href="Category.Ambient.Setoids.html#2028" class="Bound">n</a> <a id="2041" href="Agda.Builtin.Equality.html#150" class="Datatype Operator">≡</a> <a id="2043" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="2047" href="Category.Ambient.Setoids.html#2030" class="Bound">m</a> <a id="2049" class="Symbol">→</a> <a id="2051" href="Category.Ambient.Setoids.html#2028" class="Bound">n</a> <a id="2053" href="Agda.Builtin.Equality.html#150" class="Datatype Operator">≡</a> <a id="2055" href="Category.Ambient.Setoids.html#2030" class="Bound">m</a>
|
|||
|
<a id="2059" href="Category.Ambient.Setoids.html#2015" class="Function">suc-inj</a> <a id="2067" href="Agda.Builtin.Equality.html#207" class="InductiveConstructor">Eq.refl</a> <a id="2075" class="Symbol">=</a> <a id="2077" href="Agda.Builtin.Equality.html#207" class="InductiveConstructor">Eq.refl</a>
|
|||
|
|
|||
|
<a id="ℕ-eq"></a><a id="2088" href="Category.Ambient.Setoids.html#2088" class="Function">ℕ-eq</a> <a id="2093" class="Symbol">:</a> <a id="2095" href="Relation.Binary.Core.html#896" class="Function">Rel</a> <a id="2099" href="Category.Ambient.Setoids.html#1884" class="Datatype">ℕ</a> <a id="2101" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a>
|
|||
|
<a id="2105" href="Category.Ambient.Setoids.html#2088" class="Function">ℕ-eq</a> <a id="2110" href="Category.Ambient.Setoids.html#1904" class="InductiveConstructor">zero</a> <a id="2115" href="Category.Ambient.Setoids.html#1904" class="InductiveConstructor">zero</a> <a id="2120" class="Symbol">=</a> <a id="2122" href="Data.Unit.Polymorphic.Base.html#489" class="Function">⊤</a>
|
|||
|
<a id="2126" href="Category.Ambient.Setoids.html#2088" class="Function">ℕ-eq</a> <a id="2131" href="Category.Ambient.Setoids.html#1904" class="InductiveConstructor">zero</a> <a id="2136" class="Symbol">(</a><a id="2137" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="2141" href="Category.Ambient.Setoids.html#2141" class="Bound">m</a><a id="2142" class="Symbol">)</a> <a id="2144" class="Symbol">=</a> <a id="2146" href="Data.Empty.Polymorphic.html#340" class="Function">⊥</a>
|
|||
|
<a id="2150" href="Category.Ambient.Setoids.html#2088" class="Function">ℕ-eq</a> <a id="2155" class="Symbol">(</a><a id="2156" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="2160" href="Category.Ambient.Setoids.html#2160" class="Bound">n</a><a id="2161" class="Symbol">)</a> <a id="2163" href="Category.Ambient.Setoids.html#1904" class="InductiveConstructor">zero</a> <a id="2168" class="Symbol">=</a> <a id="2170" href="Data.Empty.Polymorphic.html#340" class="Function">⊥</a>
|
|||
|
<a id="2174" href="Category.Ambient.Setoids.html#2088" class="Function">ℕ-eq</a> <a id="2179" class="Symbol">(</a><a id="2180" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="2184" href="Category.Ambient.Setoids.html#2184" class="Bound">n</a><a id="2185" class="Symbol">)</a> <a id="2187" class="Symbol">(</a><a id="2188" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="2192" href="Category.Ambient.Setoids.html#2192" class="Bound">m</a><a id="2193" class="Symbol">)</a> <a id="2195" class="Symbol">=</a> <a id="2197" href="Category.Ambient.Setoids.html#2088" class="Function">ℕ-eq</a> <a id="2202" href="Category.Ambient.Setoids.html#2184" class="Bound">n</a> <a id="2204" href="Category.Ambient.Setoids.html#2192" class="Bound">m</a>
|
|||
|
|
|||
|
<a id="⊤-setoid"></a><a id="2209" href="Category.Ambient.Setoids.html#2209" class="Function">⊤-setoid</a> <a id="2218" class="Symbol">:</a> <a id="2220" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="2227" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="2229" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a>
|
|||
|
<a id="2233" href="Category.Ambient.Setoids.html#2209" class="Function">⊤-setoid</a> <a id="2242" class="Symbol">=</a> <a id="2244" class="Keyword">record</a> <a id="2251" class="Symbol">{</a> <a id="2253" href="Relation.Binary.Bundles.html#1143" class="Field">Carrier</a> <a id="2261" class="Symbol">=</a> <a id="2263" href="Data.Unit.Polymorphic.Base.html#489" class="Function">⊤</a> <a id="2265" class="Symbol">;</a> <a id="2267" href="Relation.Binary.Bundles.html#1169" class="Field Operator">_≈_</a> <a id="2271" class="Symbol">=</a> <a id="2273" href="Agda.Builtin.Equality.html#150" class="Datatype Operator">_≡_</a> <a id="2277" class="Symbol">;</a> <a id="2279" href="Relation.Binary.Bundles.html#1203" class="Field">isEquivalence</a> <a id="2293" class="Symbol">=</a> <a id="2295" href="Relation.Binary.PropositionalEquality.Properties.html#5422" class="Function">Eq.isEquivalence</a> <a id="2312" class="Symbol">}</a>
|
|||
|
|
|||
|
<a id="ℕ-setoid"></a><a id="2317" href="Category.Ambient.Setoids.html#2317" class="Function">ℕ-setoid</a> <a id="2326" class="Symbol">:</a> <a id="2328" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="2335" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="2337" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a>
|
|||
|
<a id="2341" href="Category.Ambient.Setoids.html#2317" class="Function">ℕ-setoid</a> <a id="2350" class="Symbol">=</a> <a id="2352" class="Keyword">record</a>
|
|||
|
<a id="2364" class="Symbol">{</a> <a id="2366" href="Relation.Binary.Bundles.html#1143" class="Field">Carrier</a> <a id="2374" class="Symbol">=</a> <a id="2376" href="Category.Ambient.Setoids.html#1884" class="Datatype">ℕ</a>
|
|||
|
<a id="2383" class="Symbol">;</a> <a id="2385" href="Relation.Binary.Bundles.html#1169" class="Field Operator">_≈_</a> <a id="2389" class="Symbol">=</a> <a id="2391" href="Agda.Builtin.Equality.html#150" class="Datatype Operator">_≡_</a> <a id="2395" class="Comment">-- ℕ-eq </a>
|
|||
|
<a id="2408" class="Symbol">;</a> <a id="2410" href="Relation.Binary.Bundles.html#1203" class="Field">isEquivalence</a> <a id="2424" class="Symbol">=</a> <a id="2426" href="Relation.Binary.PropositionalEquality.Properties.html#5422" class="Function">Eq.isEquivalence</a>
|
|||
|
<a id="2448" class="Symbol">}</a>
|
|||
|
|
|||
|
<a id="zero⟶"></a><a id="2455" href="Category.Ambient.Setoids.html#2455" class="Function">zero⟶</a> <a id="2461" class="Symbol">:</a> <a id="2463" href="Categories.Category.Instance.SingletonSet.html#776" class="Function">SingletonSetoid</a> <a id="2479" class="Symbol">{</a><a id="2480" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="2481" class="Symbol">}</a> <a id="2483" class="Symbol">{</a><a id="2484" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="2485" class="Symbol">}</a> <a id="2487" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="2489" href="Category.Ambient.Setoids.html#2317" class="Function">ℕ-setoid</a>
|
|||
|
<a id="2500" href="Category.Ambient.Setoids.html#2455" class="Function">zero⟶</a> <a id="2506" class="Symbol">=</a> <a id="2508" class="Keyword">record</a> <a id="2515" class="Symbol">{</a> <a id="2517" href="Function.Bundles.html#2094" class="Field">to</a> <a id="2520" class="Symbol">=</a> <a id="2522" class="Symbol">λ</a> <a id="2524" href="Category.Ambient.Setoids.html#2524" class="Bound">_</a> <a id="2526" class="Symbol">→</a> <a id="2528" href="Category.Ambient.Setoids.html#1904" class="InductiveConstructor">zero</a> <a id="2533" class="Symbol">;</a> <a id="2535" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="2540" class="Symbol">=</a> <a id="2542" class="Symbol">λ</a> <a id="2544" href="Category.Ambient.Setoids.html#2544" class="Bound">x</a> <a id="2546" class="Symbol">→</a> <a id="2548" href="Agda.Builtin.Equality.html#207" class="InductiveConstructor">Eq.refl</a> <a id="2556" class="Symbol">}</a>
|
|||
|
<a id="suc⟶"></a><a id="2561" href="Category.Ambient.Setoids.html#2561" class="Function">suc⟶</a> <a id="2566" class="Symbol">:</a> <a id="2568" href="Category.Ambient.Setoids.html#2317" class="Function">ℕ-setoid</a> <a id="2577" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="2579" href="Category.Ambient.Setoids.html#2317" class="Function">ℕ-setoid</a>
|
|||
|
<a id="2590" href="Category.Ambient.Setoids.html#2561" class="Function">suc⟶</a> <a id="2595" class="Symbol">=</a> <a id="2597" class="Keyword">record</a> <a id="2604" class="Symbol">{</a> <a id="2606" href="Function.Bundles.html#2094" class="Field">to</a> <a id="2609" class="Symbol">=</a> <a id="2611" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="2615" class="Symbol">;</a> <a id="2617" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="2622" class="Symbol">=</a> <a id="2624" href="Category.Ambient.Setoids.html#1932" class="Function">suc-cong</a> <a id="2633" class="Symbol">}</a>
|
|||
|
|
|||
|
<a id="ℕ-universal"></a><a id="2638" href="Category.Ambient.Setoids.html#2638" class="Function">ℕ-universal</a> <a id="2650" class="Symbol">:</a> <a id="2652" class="Symbol">∀</a> <a id="2654" class="Symbol">{</a><a id="2655" href="Category.Ambient.Setoids.html#2655" class="Bound">A</a> <a id="2657" class="Symbol">:</a> <a id="2659" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="2666" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="2668" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="2669" class="Symbol">}</a> <a id="2671" class="Symbol">→</a> <a id="2673" href="Categories.Category.Instance.SingletonSet.html#776" class="Function">SingletonSetoid</a> <a id="2689" class="Symbol">{</a><a id="2690" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="2691" class="Symbol">}</a> <a id="2693" class="Symbol">{</a><a id="2694" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="2695" class="Symbol">}</a> <a id="2697" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="2699" href="Category.Ambient.Setoids.html#2655" class="Bound">A</a> <a id="2701" class="Symbol">→</a> <a id="2703" href="Category.Ambient.Setoids.html#2655" class="Bound">A</a> <a id="2705" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="2707" href="Category.Ambient.Setoids.html#2655" class="Bound">A</a> <a id="2709" class="Symbol">→</a> <a id="2711" href="Category.Ambient.Setoids.html#2317" class="Function">ℕ-setoid</a> <a id="2720" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="2722" href="Category.Ambient.Setoids.html#2655" class="Bound">A</a>
|
|||
|
<a id="2726" href="Category.Ambient.Setoids.html#2638" class="Function">ℕ-universal</a> <a id="2738" class="Symbol">{</a><a id="2739" href="Category.Ambient.Setoids.html#2739" class="Bound">A</a><a id="2740" class="Symbol">}</a> <a id="2742" href="Category.Ambient.Setoids.html#2742" class="Bound">z</a> <a id="2744" href="Category.Ambient.Setoids.html#2744" class="Bound">s</a> <a id="2746" class="Symbol">=</a> <a id="2748" class="Keyword">record</a> <a id="2755" class="Symbol">{</a> <a id="2757" href="Function.Bundles.html#2094" class="Field">to</a> <a id="2760" class="Symbol">=</a> <a id="2762" href="Category.Ambient.Setoids.html#2799" class="Function">app</a> <a id="2766" class="Symbol">;</a> <a id="2768" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="2773" class="Symbol">=</a> <a id="2775" href="Category.Ambient.Setoids.html#2892" class="Function">cong'</a> <a id="2781" class="Symbol">}</a>
|
|||
|
<a id="2787" class="Keyword">where</a>
|
|||
|
<a id="2799" href="Category.Ambient.Setoids.html#2799" class="Function">app</a> <a id="2803" class="Symbol">:</a> <a id="2805" href="Category.Ambient.Setoids.html#1884" class="Datatype">ℕ</a> <a id="2807" class="Symbol">→</a> <a id="2809" href="Relation.Binary.Bundles.html#1143" class="Field">Setoid.Carrier</a> <a id="2824" href="Category.Ambient.Setoids.html#2739" class="Bound">A</a>
|
|||
|
<a id="2832" href="Category.Ambient.Setoids.html#2799" class="Function">app</a> <a id="2836" href="Category.Ambient.Setoids.html#1904" class="InductiveConstructor">zero</a> <a id="2841" class="Symbol">=</a> <a id="2843" href="Category.Ambient.Setoids.html#2742" class="Bound">z</a> <a id="2845" href="Function.Bundles.html#15133" class="Function Operator">⟨$⟩</a> <a id="2849" href="Data.Unit.Polymorphic.Base.html#533" class="Function">tt</a>
|
|||
|
<a id="2858" href="Category.Ambient.Setoids.html#2799" class="Function">app</a> <a id="2862" class="Symbol">(</a><a id="2863" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="2867" href="Category.Ambient.Setoids.html#2867" class="Bound">n</a><a id="2868" class="Symbol">)</a> <a id="2870" class="Symbol">=</a> <a id="2872" href="Category.Ambient.Setoids.html#2744" class="Bound">s</a> <a id="2874" href="Function.Bundles.html#15133" class="Function Operator">⟨$⟩</a> <a id="2878" class="Symbol">(</a><a id="2879" href="Category.Ambient.Setoids.html#2799" class="Function">app</a> <a id="2883" href="Category.Ambient.Setoids.html#2867" class="Bound">n</a><a id="2884" class="Symbol">)</a>
|
|||
|
<a id="2892" href="Category.Ambient.Setoids.html#2892" class="Function">cong'</a> <a id="2898" class="Symbol">:</a> <a id="2900" class="Symbol">∀</a> <a id="2902" class="Symbol">{</a><a id="2903" href="Category.Ambient.Setoids.html#2903" class="Bound">n</a> <a id="2905" href="Category.Ambient.Setoids.html#2905" class="Bound">m</a> <a id="2907" class="Symbol">:</a> <a id="2909" href="Category.Ambient.Setoids.html#1884" class="Datatype">ℕ</a><a id="2910" class="Symbol">}</a> <a id="2912" class="Symbol">→</a> <a id="2914" href="Category.Ambient.Setoids.html#2903" class="Bound">n</a> <a id="2916" href="Agda.Builtin.Equality.html#150" class="Datatype Operator">≡</a> <a id="2918" href="Category.Ambient.Setoids.html#2905" class="Bound">m</a> <a id="2920" class="Symbol">→</a> <a id="2922" href="Relation.Binary.Bundles.html#1169" class="Field Operator">Setoid._≈_</a> <a id="2933" href="Category.Ambient.Setoids.html#2739" class="Bound">A</a> <a id="2935" class="Symbol">(</a><a id="2936" href="Category.Ambient.Setoids.html#2799" class="Function">app</a> <a id="2940" href="Category.Ambient.Setoids.html#2903" class="Bound">n</a><a id="2941" class="Symbol">)</a> <a id="2943" class="Symbol">(</a><a id="2944" href="Category.Ambient.Setoids.html#2799" class="Function">app</a> <a id="2948" href="Category.Ambient.Setoids.html#2905" class="Bound">m</a><a id="2949" class="Symbol">)</a>
|
|||
|
<a id="2957" href="Category.Ambient.Setoids.html#2892" class="Function">cong'</a> <a id="2963" href="Agda.Builtin.Equality.html#207" class="InductiveConstructor">Eq.refl</a> <a id="2971" class="Symbol">=</a> <a id="2973" href="Relation.Binary.Structures.html#1596" class="Field">IsEquivalence.refl</a> <a id="2992" class="Symbol">(</a><a id="2993" href="Relation.Binary.Bundles.html#1203" class="Field">Setoid.isEquivalence</a> <a id="3014" href="Category.Ambient.Setoids.html#2739" class="Bound">A</a><a id="3015" class="Symbol">)</a>
|
|||
|
|
|||
|
<a id="ℕ-z-commute"></a><a id="3024" href="Category.Ambient.Setoids.html#3024" class="Function">ℕ-z-commute</a> <a id="3036" class="Symbol">:</a> <a id="3038" class="Symbol">∀</a> <a id="3040" class="Symbol">{</a><a id="3041" href="Category.Ambient.Setoids.html#3041" class="Bound">A</a> <a id="3043" class="Symbol">:</a> <a id="3045" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="3052" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="3054" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="3055" class="Symbol">}</a> <a id="3057" class="Symbol">{</a><a id="3058" href="Category.Ambient.Setoids.html#3058" class="Bound">q</a> <a id="3060" class="Symbol">:</a> <a id="3062" href="Categories.Category.Instance.SingletonSet.html#776" class="Function">SingletonSetoid</a> <a id="3078" class="Symbol">{</a><a id="3079" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="3080" class="Symbol">}</a> <a id="3082" class="Symbol">{</a><a id="3083" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="3084" class="Symbol">}</a> <a id="3086" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="3088" href="Category.Ambient.Setoids.html#3041" class="Bound">A</a><a id="3089" class="Symbol">}</a> <a id="3091" class="Symbol">{</a><a id="3092" href="Category.Ambient.Setoids.html#3092" class="Bound">f</a> <a id="3094" class="Symbol">:</a> <a id="3096" href="Category.Ambient.Setoids.html#3041" class="Bound">A</a> <a id="3098" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="3100" href="Category.Ambient.Setoids.html#3041" class="Bound">A</a><a id="3101" class="Symbol">}</a> <a id="3103" class="Symbol">→</a> <a id="3105" href="Category.Ambient.Setoids.html#3058" class="Bound">q</a> <a id="3107" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="3109" class="Symbol">(</a><a id="3110" href="Category.Ambient.Setoids.html#2638" class="Function">ℕ-universal</a> <a id="3122" href="Category.Ambient.Setoids.html#3058" class="Bound">q</a> <a id="3124" href="Category.Ambient.Setoids.html#3092" class="Bound">f</a> <a id="3126" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3128" href="Category.Ambient.Setoids.html#2455" class="Function">zero⟶</a><a id="3133" class="Symbol">)</a>
|
|||
|
<a id="3137" href="Category.Ambient.Setoids.html#3024" class="Function">ℕ-z-commute</a> <a id="3149" class="Symbol">{</a><a id="3150" href="Category.Ambient.Setoids.html#3150" class="Bound">A</a><a id="3151" class="Symbol">}</a> <a id="3153" class="Symbol">{</a><a id="3154" href="Category.Ambient.Setoids.html#3154" class="Bound">q</a><a id="3155" class="Symbol">}</a> <a id="3157" class="Symbol">{</a><a id="3158" href="Category.Ambient.Setoids.html#3158" class="Bound">f</a><a id="3159" class="Symbol">}</a> <a id="3161" class="Symbol">{</a><a id="3162" href="Level.html#466" class="InductiveConstructor">lift</a> <a id="3167" href="Category.Ambient.Setoids.html#3167" class="Bound">t</a><a id="3168" class="Symbol">}</a> <a id="3170" class="Symbol">=</a> <a id="3172" href="Relation.Binary.Structures.html#1596" class="Field">IsEquivalence.refl</a> <a id="3191" class="Symbol">(</a><a id="3192" href="Relation.Binary.Bundles.html#1203" class="Field">Setoid.isEquivalence</a> <a id="3213" href="Category.Ambient.Setoids.html#3150" class="Bound">A</a><a id="3214" class="Symbol">)</a>
|
|||
|
|
|||
|
<a id="ℕ-s-commute"></a><a id="3219" href="Category.Ambient.Setoids.html#3219" class="Function">ℕ-s-commute</a> <a id="3231" class="Symbol">:</a> <a id="3233" class="Symbol">∀</a> <a id="3235" class="Symbol">{</a><a id="3236" href="Category.Ambient.Setoids.html#3236" class="Bound">A</a> <a id="3238" class="Symbol">:</a> <a id="3240" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="3247" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="3249" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="3250" class="Symbol">}</a> <a id="3252" class="Symbol">{</a><a id="3253" href="Category.Ambient.Setoids.html#3253" class="Bound">q</a> <a id="3255" class="Symbol">:</a> <a id="3257" href="Categories.Category.Instance.SingletonSet.html#776" class="Function">SingletonSetoid</a> <a id="3273" class="Symbol">{</a><a id="3274" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="3275" class="Symbol">}</a> <a id="3277" class="Symbol">{</a><a id="3278" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="3279" class="Symbol">}</a> <a id="3281" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="3283" href="Category.Ambient.Setoids.html#3236" class="Bound">A</a><a id="3284" class="Symbol">}</a> <a id="3286" class="Symbol">{</a><a id="3287" href="Category.Ambient.Setoids.html#3287" class="Bound">f</a> <a id="3289" class="Symbol">:</a> <a id="3291" href="Category.Ambient.Setoids.html#3236" class="Bound">A</a> <a id="3293" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="3295" href="Category.Ambient.Setoids.html#3236" class="Bound">A</a><a id="3296" class="Symbol">}</a> <a id="3298" class="Symbol">→</a> <a id="3300" class="Symbol">(</a><a id="3301" href="Category.Ambient.Setoids.html#3287" class="Bound">f</a> <a id="3303" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3305" class="Symbol">(</a><a id="3306" href="Category.Ambient.Setoids.html#2638" class="Function">ℕ-universal</a> <a id="3318" href="Category.Ambient.Setoids.html#3253" class="Bound">q</a> <a id="3320" href="Category.Ambient.Setoids.html#3287" class="Bound">f</a><a id="3321" class="Symbol">))</a> <a id="3324" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="3326" class="Symbol">(</a><a id="3327" href="Category.Ambient.Setoids.html#2638" class="Function">ℕ-universal</a> <a id="3339" href="Category.Ambient.Setoids.html#3253" class="Bound">q</a> <a id="3341" href="Category.Ambient.Setoids.html#3287" class="Bound">f</a> <a id="3343" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3345" href="Category.Ambient.Setoids.html#2561" class="Function">suc⟶</a><a id="3349" class="Symbol">)</a>
|
|||
|
<a id="3353" href="Category.Ambient.Setoids.html#3219" class="Function">ℕ-s-commute</a> <a id="3365" class="Symbol">{</a><a id="3366" href="Category.Ambient.Setoids.html#3366" class="Bound">A</a><a id="3367" class="Symbol">}</a> <a id="3369" class="Symbol">{</a><a id="3370" href="Category.Ambient.Setoids.html#3370" class="Bound">q</a><a id="3371" class="Symbol">}</a> <a id="3373" class="Symbol">{</a><a id="3374" href="Category.Ambient.Setoids.html#3374" class="Bound">f</a><a id="3375" class="Symbol">}</a> <a id="3377" class="Symbol">{</a><a id="3378" href="Category.Ambient.Setoids.html#3378" class="Bound">n</a><a id="3379" class="Symbol">}</a> <a id="3381" class="Symbol">=</a> <a id="3383" href="Relation.Binary.Structures.html#1596" class="Field">IsEquivalence.refl</a> <a id="3402" class="Symbol">(</a><a id="3403" href="Relation.Binary.Bundles.html#1203" class="Field">Setoid.isEquivalence</a> <a id="3424" href="Category.Ambient.Setoids.html#3366" class="Bound">A</a><a id="3425" class="Symbol">)</a>
|
|||
|
|
|||
|
<a id="ℕ-unique"></a><a id="3430" href="Category.Ambient.Setoids.html#3430" class="Function">ℕ-unique</a> <a id="3439" class="Symbol">:</a> <a id="3441" class="Symbol">∀</a> <a id="3443" class="Symbol">{</a><a id="3444" href="Category.Ambient.Setoids.html#3444" class="Bound">A</a> <a id="3446" class="Symbol">:</a> <a id="3448" href="Relation.Binary.Bundles.html#1080" class="Record">Setoid</a> <a id="3455" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="3457" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="3458" class="Symbol">}</a> <a id="3460" class="Symbol">{</a><a id="3461" href="Category.Ambient.Setoids.html#3461" class="Bound">q</a> <a id="3463" class="Symbol">:</a> <a id="3465" href="Categories.Category.Instance.SingletonSet.html#776" class="Function">SingletonSetoid</a> <a id="3481" class="Symbol">{</a><a id="3482" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="3483" class="Symbol">}</a> <a id="3485" class="Symbol">{</a><a id="3486" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="3487" class="Symbol">}</a> <a id="3489" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="3491" href="Category.Ambient.Setoids.html#3444" class="Bound">A</a><a id="3492" class="Symbol">}</a> <a id="3494" class="Symbol">{</a><a id="3495" href="Category.Ambient.Setoids.html#3495" class="Bound">f</a> <a id="3497" class="Symbol">:</a> <a id="3499" href="Category.Ambient.Setoids.html#3444" class="Bound">A</a> <a id="3501" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="3503" href="Category.Ambient.Setoids.html#3444" class="Bound">A</a><a id="3504" class="Symbol">}</a> <a id="3506" class="Symbol">{</a><a id="3507" href="Category.Ambient.Setoids.html#3507" class="Bound">u</a> <a id="3509" class="Symbol">:</a> <a id="3511" href="Category.Ambient.Setoids.html#2317" class="Function">ℕ-setoid</a> <a id="3520" href="Category.Ambient.Setoids.html#626" class="Record Operator">⟶</a> <a id="3522" href="Category.Ambient.Setoids.html#3444" class="Bound">A</a><a id="3523" class="Symbol">}</a> <a id="3525" class="Symbol">→</a> <a id="3527" href="Category.Ambient.Setoids.html#3461" class="Bound">q</a> <a id="3529" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="3531" class="Symbol">(</a><a id="3532" href="Category.Ambient.Setoids.html#3507" class="Bound">u</a> <a id="3534" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3536" href="Category.Ambient.Setoids.html#2455" class="Function">zero⟶</a><a id="3541" class="Symbol">)</a> <a id="3543" class="Symbol">→</a> <a id="3545" class="Symbol">(</a><a id="3546" href="Category.Ambient.Setoids.html#3495" class="Bound">f</a> <a id="3548" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3550" href="Category.Ambient.Setoids.html#3507" class="Bound">u</a><a id="3551" class="Symbol">)</a> <a id="3553" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="3555" class="Symbol">(</a><a id="3556" href="Category.Ambient.Setoids.html#3507" class="Bound">u</a> <a id="3558" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3560" href="Category.Ambient.Setoids.html#2561" class="Function">suc⟶</a><a id="3564" class="Symbol">)</a> <a id="3566" class="Symbol">→</a> <a id="3568" href="Category.Ambient.Setoids.html#3507" class="Bound">u</a> <a id="3570" href="Category.Ambient.Setoids.html#1335" class="Function Operator">≋</a> <a id="3572" href="Category.Ambient.Setoids.html#2638" class="Function">ℕ-universal</a> <a id="3584" href="Category.Ambient.Setoids.html#3461" class="Bound">q</a> <a id="3586" href="Category.Ambient.Setoids.html#3495" class="Bound">f</a>
|
|||
|
<a id="3590" href="Category.Ambient.Setoids.html#3430" class="Function">ℕ-unique</a> <a id="3599" class="Symbol">{</a><a id="3600" href="Category.Ambient.Setoids.html#3600" class="Bound">A</a><a id="3601" class="Symbol">}</a> <a id="3603" class="Symbol">{</a><a id="3604" href="Category.Ambient.Setoids.html#3604" class="Bound">q</a><a id="3605" class="Symbol">}</a> <a id="3607" class="Symbol">{</a><a id="3608" href="Category.Ambient.Setoids.html#3608" class="Bound">f</a><a id="3609" class="Symbol">}</a> <a id="3611" class="Symbol">{</a><a id="3612" href="Category.Ambient.Setoids.html#3612" class="Bound">u</a><a id="3613" class="Symbol">}</a> <a id="3615" href="Category.Ambient.Setoids.html#3615" class="Bound">qz</a> <a id="3618" href="Category.Ambient.Setoids.html#3618" class="Bound">fs</a> <a id="3621" class="Symbol">{</a><a id="3622" href="Category.Ambient.Setoids.html#1904" class="InductiveConstructor">zero</a><a id="3626" class="Symbol">}</a> <a id="3628" class="Symbol">=</a> <a id="3630" href="Category.Ambient.Setoids.html#1435" class="Function">≋-sym</a> <a id="3636" class="Symbol">{</a><a id="3637" href="Categories.Category.Instance.SingletonSet.html#776" class="Function">SingletonSetoid</a><a id="3652" class="Symbol">}</a> <a id="3654" class="Symbol">{</a><a id="3655" href="Category.Ambient.Setoids.html#3600" class="Bound">A</a><a id="3656" class="Symbol">}</a> <a id="3658" class="Symbol">{</a><a id="3659" href="Category.Ambient.Setoids.html#3604" class="Bound">q</a><a id="3660" class="Symbol">}</a> <a id="3662" class="Symbol">{</a><a id="3663" href="Category.Ambient.Setoids.html#3612" class="Bound">u</a> <a id="3665" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3667" href="Category.Ambient.Setoids.html#2455" class="Function">zero⟶</a><a id="3672" class="Symbol">}</a> <a id="3674" href="Category.Ambient.Setoids.html#3615" class="Bound">qz</a>
|
|||
|
<a id="3679" href="Category.Ambient.Setoids.html#3430" class="Function">ℕ-unique</a> <a id="3688" class="Symbol">{</a><a id="3689" href="Category.Ambient.Setoids.html#3689" class="Bound">A</a><a id="3690" class="Symbol">}</a> <a id="3692" class="Symbol">{</a><a id="3693" href="Category.Ambient.Setoids.html#3693" class="Bound">q</a><a id="3694" class="Symbol">}</a> <a id="3696" class="Symbol">{</a><a id="3697" href="Category.Ambient.Setoids.html#3697" class="Bound">f</a><a id="3698" class="Symbol">}</a> <a id="3700" class="Symbol">{</a><a id="3701" href="Category.Ambient.Setoids.html#3701" class="Bound">u</a><a id="3702" class="Symbol">}</a> <a id="3704" href="Category.Ambient.Setoids.html#3704" class="Bound">qz</a> <a id="3707" href="Category.Ambient.Setoids.html#3707" class="Bound">fs</a> <a id="3710" class="Symbol">{</a><a id="3711" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="3715" href="Category.Ambient.Setoids.html#3715" class="Bound">n</a><a id="3716" class="Symbol">}</a> <a id="3718" class="Symbol">=</a> <a id="3720" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">AR.begin</a>
|
|||
|
<a id="3734" href="Category.Ambient.Setoids.html#3701" class="Bound">u</a> <a id="3736" href="Function.Bundles.html#15133" class="Function Operator">⟨$⟩</a> <a id="3740" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="3744" href="Category.Ambient.Setoids.html#3715" class="Bound">n</a> <a id="3789" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">AR.≈⟨</a> <a id="3795" href="Category.Ambient.Setoids.html#1435" class="Function">≋-sym</a> <a id="3801" class="Symbol">{</a><a id="3802" href="Category.Ambient.Setoids.html#2317" class="Function">ℕ-setoid</a><a id="3810" class="Symbol">}</a> <a id="3812" class="Symbol">{</a><a id="3813" href="Category.Ambient.Setoids.html#3689" class="Bound">A</a><a id="3814" class="Symbol">}</a> <a id="3816" class="Symbol">{</a><a id="3817" href="Category.Ambient.Setoids.html#3697" class="Bound">f</a> <a id="3819" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3821" href="Category.Ambient.Setoids.html#3701" class="Bound">u</a><a id="3822" class="Symbol">}</a> <a id="3824" class="Symbol">{</a><a id="3825" href="Category.Ambient.Setoids.html#3701" class="Bound">u</a> <a id="3827" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3829" href="Category.Ambient.Setoids.html#2561" class="Function">suc⟶</a><a id="3833" class="Symbol">}</a> <a id="3835" href="Category.Ambient.Setoids.html#3707" class="Bound">fs</a> <a id="3838" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
|||
|
<a id="3845" href="Category.Ambient.Setoids.html#3697" class="Bound">f</a> <a id="3847" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3849" href="Category.Ambient.Setoids.html#3701" class="Bound">u</a> <a id="3851" href="Function.Bundles.html#15133" class="Function Operator">⟨$⟩</a> <a id="3855" href="Category.Ambient.Setoids.html#3715" class="Bound">n</a> <a id="3900" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">AR.≈⟨</a> <a id="3906" href="Function.Bundles.html#2113" class="Field">cong</a> <a id="3911" href="Category.Ambient.Setoids.html#3697" class="Bound">f</a> <a id="3913" class="Symbol">(</a><a id="3914" href="Category.Ambient.Setoids.html#3430" class="Function">ℕ-unique</a> <a id="3923" class="Symbol">{</a><a id="3924" href="Category.Ambient.Setoids.html#3689" class="Bound">A</a><a id="3925" class="Symbol">}</a> <a id="3927" class="Symbol">{</a><a id="3928" href="Category.Ambient.Setoids.html#3693" class="Bound">q</a><a id="3929" class="Symbol">}</a> <a id="3931" class="Symbol">{</a><a id="3932" href="Category.Ambient.Setoids.html#3697" class="Bound">f</a><a id="3933" class="Symbol">}</a> <a id="3935" class="Symbol">{</a><a id="3936" href="Category.Ambient.Setoids.html#3701" class="Bound">u</a><a id="3937" class="Symbol">}</a> <a id="3939" href="Category.Ambient.Setoids.html#3704" class="Bound">qz</a> <a id="3942" href="Category.Ambient.Setoids.html#3707" class="Bound">fs</a> <a id="3945" class="Symbol">{</a><a id="3946" href="Category.Ambient.Setoids.html#3715" class="Bound">n</a><a id="3947" class="Symbol">})</a> <a id="3950" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
|||
|
<a id="3957" href="Category.Ambient.Setoids.html#3697" class="Bound">f</a> <a id="3959" href="Category.Ambient.Setoids.html#748" class="Function Operator">∘</a> <a id="3961" href="Category.Ambient.Setoids.html#2638" class="Function">ℕ-universal</a> <a id="3973" href="Category.Ambient.Setoids.html#3693" class="Bound">q</a> <a id="3975" href="Category.Ambient.Setoids.html#3697" class="Bound">f</a> <a id="3977" href="Function.Bundles.html#15133" class="Function Operator">⟨$⟩</a> <a id="3981" href="Category.Ambient.Setoids.html#3715" class="Bound">n</a> <a id="4012" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">AR.≈⟨</a> <a id="4018" href="Category.Ambient.Setoids.html#3219" class="Function">ℕ-s-commute</a> <a id="4030" class="Symbol">{</a><a id="4031" href="Category.Ambient.Setoids.html#3689" class="Bound">A</a><a id="4032" class="Symbol">}</a> <a id="4034" class="Symbol">{</a><a id="4035" href="Category.Ambient.Setoids.html#3693" class="Bound">q</a><a id="4036" class="Symbol">}</a> <a id="4038" class="Symbol">{</a><a id="4039" href="Category.Ambient.Setoids.html#3697" class="Bound">f</a><a id="4040" class="Symbol">}</a> <a id="4042" class="Symbol">{</a><a id="4043" href="Category.Ambient.Setoids.html#3715" class="Bound">n</a><a id="4044" class="Symbol">}</a> <a id="4046" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
|||
|
<a id="4053" href="Category.Ambient.Setoids.html#2638" class="Function">ℕ-universal</a> <a id="4065" href="Category.Ambient.Setoids.html#3693" class="Bound">q</a> <a id="4067" href="Category.Ambient.Setoids.html#3697" class="Bound">f</a> <a id="4069" href="Function.Bundles.html#15133" class="Function Operator">⟨$⟩</a> <a id="4073" href="Category.Ambient.Setoids.html#1917" class="InductiveConstructor">suc</a> <a id="4077" href="Category.Ambient.Setoids.html#3715" class="Bound">n</a> <a id="4108" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">AR.∎</a>
|
|||
|
<a id="4117" class="Keyword">where</a> <a id="4123" class="Keyword">module</a> <a id="4130" href="Category.Ambient.Setoids.html#4130" class="Module">AR</a> <a id="4133" class="Symbol">=</a> <a id="4135" href="Relation.Binary.Reasoning.Setoid.html" class="Module">SetoidR</a> <a id="4143" href="Category.Ambient.Setoids.html#3689" class="Bound">A</a>
|
|||
|
|
|||
|
<a id="setoidNNO"></a><a id="4148" href="Category.Ambient.Setoids.html#4148" class="Function">setoidNNO</a> <a id="4158" class="Symbol">:</a> <a id="4160" href="Categories.Object.NaturalNumbers.html#1072" class="Record">NNO</a> <a id="4164" class="Symbol">(</a><a id="4165" href="Categories.Category.Instance.Setoids.html#555" class="Function">Setoids</a> <a id="4173" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="4175" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="4176" class="Symbol">)</a> <a id="4178" href="Categories.Category.Instance.SingletonSet.html#869" class="Function">SingletonSetoid-⊤</a>
|
|||
|
<a id="4198" href="Category.Ambient.Setoids.html#4148" class="Function">setoidNNO</a> <a id="4208" class="Symbol">=</a> <a id="4210" class="Keyword">record</a>
|
|||
|
<a id="4222" class="Symbol">{</a> <a id="4224" href="Categories.Object.NaturalNumbers.html#1112" class="Field">N</a> <a id="4226" class="Symbol">=</a> <a id="4228" href="Category.Ambient.Setoids.html#2317" class="Function">ℕ-setoid</a>
|
|||
|
<a id="4242" class="Symbol">;</a> <a id="4244" href="Categories.Object.NaturalNumbers.html#1124" class="Field">isNNO</a> <a id="4250" class="Symbol">=</a> <a id="4252" class="Keyword">record</a>
|
|||
|
<a id="4265" class="Symbol">{</a> <a id="4267" href="Categories.Object.NaturalNumbers.html#464" class="Field">z</a> <a id="4269" class="Symbol">=</a> <a id="4271" href="Category.Ambient.Setoids.html#2455" class="Function">zero⟶</a>
|
|||
|
<a id="4283" class="Symbol">;</a> <a id="4285" href="Categories.Object.NaturalNumbers.html#478" class="Field">s</a> <a id="4287" class="Symbol">=</a> <a id="4289" href="Category.Ambient.Setoids.html#2561" class="Function">suc⟶</a>
|
|||
|
<a id="4300" class="Symbol">;</a> <a id="4302" href="Categories.Object.NaturalNumbers.html#492" class="Field">universal</a> <a id="4312" class="Symbol">=</a> <a id="4314" href="Category.Ambient.Setoids.html#2638" class="Function">ℕ-universal</a>
|
|||
|
<a id="4332" class="Symbol">;</a> <a id="4334" href="Categories.Object.NaturalNumbers.html#538" class="Field">z-commute</a> <a id="4344" class="Symbol">=</a> <a id="4346" class="Symbol">λ</a> <a id="4348" class="Symbol">{</a><a id="4349" href="Category.Ambient.Setoids.html#4349" class="Bound">A</a><a id="4350" class="Symbol">}</a> <a id="4352" class="Symbol">{</a><a id="4353" href="Category.Ambient.Setoids.html#4353" class="Bound">q</a><a id="4354" class="Symbol">}</a> <a id="4356" class="Symbol">{</a><a id="4357" href="Category.Ambient.Setoids.html#4357" class="Bound">f</a><a id="4358" class="Symbol">}</a> <a id="4360" class="Symbol">→</a> <a id="4362" href="Category.Ambient.Setoids.html#3024" class="Function">ℕ-z-commute</a> <a id="4374" class="Symbol">{</a><a id="4375" href="Category.Ambient.Setoids.html#4349" class="Bound">A</a><a id="4376" class="Symbol">}</a> <a id="4378" class="Symbol">{</a><a id="4379" href="Category.Ambient.Setoids.html#4353" class="Bound">q</a><a id="4380" class="Symbol">}</a> <a id="4382" class="Symbol">{</a><a id="4383" href="Category.Ambient.Setoids.html#4357" class="Bound">f</a><a id="4384" class="Symbol">}</a>
|
|||
|
<a id="4392" class="Symbol">;</a> <a id="4394" href="Categories.Object.NaturalNumbers.html#608" class="Field">s-commute</a> <a id="4404" class="Symbol">=</a> <a id="4406" class="Symbol">λ</a> <a id="4408" class="Symbol">{</a><a id="4409" href="Category.Ambient.Setoids.html#4409" class="Bound">A</a><a id="4410" class="Symbol">}</a> <a id="4412" class="Symbol">{</a><a id="4413" href="Category.Ambient.Setoids.html#4413" class="Bound">q</a><a id="4414" class="Symbol">}</a> <a id="4416" class="Symbol">{</a><a id="4417" href="Category.Ambient.Setoids.html#4417" class="Bound">f</a><a id="4418" class="Symbol">}</a> <a id="4420" class="Symbol">{</a><a id="4421" href="Category.Ambient.Setoids.html#4421" class="Bound">n</a><a id="4422" class="Symbol">}</a> <a id="4424" class="Symbol">→</a> <a id="4426" href="Category.Ambient.Setoids.html#3219" class="Function">ℕ-s-commute</a> <a id="4438" class="Symbol">{</a><a id="4439" href="Category.Ambient.Setoids.html#4409" class="Bound">A</a><a id="4440" class="Symbol">}</a> <a id="4442" class="Symbol">{</a><a id="4443" href="Category.Ambient.Setoids.html#4413" class="Bound">q</a><a id="4444" class="Symbol">}</a> <a id="4446" class="Symbol">{</a><a id="4447" href="Category.Ambient.Setoids.html#4417" class="Bound">f</a><a id="4448" class="Symbol">}</a> <a id="4450" class="Symbol">{</a><a id="4451" href="Category.Ambient.Setoids.html#4421" class="Bound">n</a><a id="4452" class="Symbol">}</a>
|
|||
|
<a id="4460" class="Symbol">;</a> <a id="4462" href="Categories.Object.NaturalNumbers.html#694" class="Field">unique</a> <a id="4469" class="Symbol">=</a> <a id="4471" class="Symbol">λ</a> <a id="4473" class="Symbol">{</a><a id="4474" href="Category.Ambient.Setoids.html#4474" class="Bound">A</a><a id="4475" class="Symbol">}</a> <a id="4477" class="Symbol">{</a><a id="4478" href="Category.Ambient.Setoids.html#4478" class="Bound">q</a><a id="4479" class="Symbol">}</a> <a id="4481" class="Symbol">{</a><a id="4482" href="Category.Ambient.Setoids.html#4482" class="Bound">f</a><a id="4483" class="Symbol">}</a> <a id="4485" class="Symbol">{</a><a id="4486" href="Category.Ambient.Setoids.html#4486" class="Bound">u</a><a id="4487" class="Symbol">}</a> <a id="4489" href="Category.Ambient.Setoids.html#4489" class="Bound">qz</a> <a id="4492" href="Category.Ambient.Setoids.html#4492" class="Bound">fs</a> <a id="4495" class="Symbol">→</a> <a id="4497" href="Category.Ambient.Setoids.html#3430" class="Function">ℕ-unique</a> <a id="4506" class="Symbol">{</a><a id="4507" href="Category.Ambient.Setoids.html#4474" class="Bound">A</a><a id="4508" class="Symbol">}</a> <a id="4510" class="Symbol">{</a><a id="4511" href="Category.Ambient.Setoids.html#4478" class="Bound">q</a><a id="4512" class="Symbol">}</a> <a id="4514" class="Symbol">{</a><a id="4515" href="Category.Ambient.Setoids.html#4482" class="Bound">f</a><a id="4516" class="Symbol">}</a> <a id="4518" class="Symbol">{</a><a id="4519" href="Category.Ambient.Setoids.html#4486" class="Bound">u</a><a id="4520" class="Symbol">}</a> <a id="4522" href="Category.Ambient.Setoids.html#4489" class="Bound">qz</a> <a id="4525" href="Category.Ambient.Setoids.html#4492" class="Bound">fs</a>
|
|||
|
<a id="4534" class="Symbol">}</a>
|
|||
|
<a id="4540" class="Symbol">}</a>
|
|||
|
|
|||
|
<a id="setoidAmbient"></a><a id="4545" href="Category.Ambient.Setoids.html#4545" class="Function">setoidAmbient</a> <a id="4559" class="Symbol">:</a> <a id="4561" href="Category.Ambient.html#1681" class="Record">Ambient</a> <a id="4569" class="Symbol">(</a><a id="4570" href="Category.Ambient.Setoids.html#68" class="Primitive">ℓ-suc</a> <a id="4576" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="4577" class="Symbol">)</a> <a id="4579" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="4581" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a>
|
|||
|
<a id="4585" href="Category.Ambient.Setoids.html#4545" class="Function">setoidAmbient</a> <a id="4599" class="Symbol">=</a> <a id="4601" class="Keyword">record</a> <a id="4608" class="Symbol">{</a> <a id="4610" href="Category.Ambient.html#1757" class="Field">C</a> <a id="4612" class="Symbol">=</a> <a id="4614" href="Categories.Category.Instance.Setoids.html#555" class="Function">Setoids</a> <a id="4622" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="4624" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="4626" class="Symbol">;</a> <a id="4628" href="Category.Ambient.html#1782" class="Field">extensive</a> <a id="4638" class="Symbol">=</a> <a id="4640" href="Categories.Category.Instance.Properties.Setoids.Extensive.html#1300" class="Function">Setoids-Extensive</a> <a id="4658" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="4660" class="Symbol">;</a> <a id="4662" href="Category.Ambient.html#1812" class="Field">cartesian</a> <a id="4672" class="Symbol">=</a> <a id="4674" href="Categories.Category.Monoidal.Instance.Setoids.html#1204" class="Function">Setoids-Cartesian</a> <a id="4692" class="Symbol">;</a> <a id="4694" href="Category.Ambient.html#1842" class="Field">ℕ</a> <a id="4696" class="Symbol">=</a> <a id="4698" href="Categories.Object.NaturalNumbers.Properties.Parametrized.html#1305" class="Function">NNO×CCC⇒PNNO</a> <a id="4711" class="Symbol">(</a><a id="4712" class="Keyword">record</a> <a id="4719" class="Symbol">{</a> <a id="4721" href="Categories.Category.CartesianClosed.Bundle.html#418" class="Field">U</a> <a id="4723" class="Symbol">=</a> <a id="4725" href="Categories.Category.Instance.Setoids.html#555" class="Function">Setoids</a> <a id="4733" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="4735" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="4737" class="Symbol">;</a> <a id="4739" href="Categories.Category.CartesianClosed.Bundle.html#470" class="Field">cartesianClosed</a> <a id="4755" class="Symbol">=</a> <a id="4757" href="Categories.Category.Instance.Properties.Setoids.CCC.html#2117" class="Function">Setoids-CCC</a> <a id="4769" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a> <a id="4771" class="Symbol">})</a> <a id="4774" class="Symbol">(</a><a id="4775" href="Categories.Category.Cocartesian.html#3462" class="Field">Cocartesian.coproducts</a> <a id="4798" class="Symbol">(</a><a id="4799" href="Categories.Category.Monoidal.Instance.Setoids.html#1968" class="Function">Setoids-Cocartesian</a> <a id="4819" class="Symbol">{</a><a id="4820" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="4821" class="Symbol">}</a> <a id="4823" class="Symbol">{</a><a id="4824" href="Category.Ambient.Setoids.html#1253" class="Bound">ℓ</a><a id="4825" class="Symbol">}))</a> <a id="4829" href="Category.Ambient.Setoids.html#4148" class="Function">setoidNNO</a> <a id="4839" class="Symbol">}</a>
|
|||
|
</pre>
|
|||
|
</body>
|
|||
|
</html>
|