mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
211 lines
No EOL
47 KiB
HTML
211 lines
No EOL
47 KiB
HTML
<!DOCTYPE HTML>
|
|
<html><head><meta charset="utf-8"><title>Algebra.Morphism</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">-- Morphisms between algebraic structures</a>
|
|
<a id="148" class="Comment">------------------------------------------------------------------------</a>
|
|
|
|
<a id="222" class="Symbol">{-#</a> <a id="226" class="Keyword">OPTIONS</a> <a id="234" class="Pragma">--cubical-compatible</a> <a id="255" class="Pragma">--safe</a> <a id="262" class="Symbol">#-}</a>
|
|
|
|
<a id="267" class="Keyword">module</a> <a id="274" href="Algebra.Morphism.html" class="Module">Algebra.Morphism</a> <a id="291" class="Keyword">where</a>
|
|
|
|
<a id="298" class="Keyword">import</a> <a id="305" href="Algebra.Morphism.Definitions.html" class="Module">Algebra.Morphism.Definitions</a> <a id="334" class="Symbol">as</a> <a id="337" class="Module">MorphismDefinitions</a>
|
|
<a id="357" class="Keyword">open</a> <a id="362" class="Keyword">import</a> <a id="369" href="Algebra.html" class="Module">Algebra</a>
|
|
<a id="377" class="Keyword">import</a> <a id="384" href="Algebra.Properties.Group.html" class="Module">Algebra.Properties.Group</a> <a id="409" class="Symbol">as</a> <a id="412" class="Module">GroupP</a>
|
|
<a id="419" class="Keyword">open</a> <a id="424" class="Keyword">import</a> <a id="431" href="Function.Base.html" class="Module">Function.Base</a>
|
|
<a id="445" class="Keyword">open</a> <a id="450" class="Keyword">import</a> <a id="457" href="Level.html" class="Module">Level</a>
|
|
<a id="463" class="Keyword">open</a> <a id="468" class="Keyword">import</a> <a id="475" href="Relation.Binary.Core.html" class="Module">Relation.Binary.Core</a> <a id="496" class="Keyword">using</a> <a id="502" class="Symbol">(</a><a id="503" href="Relation.Binary.Core.html#896" class="Function">Rel</a><a id="506" class="Symbol">;</a> <a id="508" href="Relation.Binary.Core.html#1577" class="Function Operator">_Preserves_⟶_</a><a id="521" class="Symbol">)</a>
|
|
<a id="523" class="Keyword">import</a> <a id="530" href="Relation.Binary.Reasoning.Setoid.html" class="Module">Relation.Binary.Reasoning.Setoid</a> <a id="563" class="Symbol">as</a> <a id="566" class="Module">EqR</a>
|
|
|
|
<a id="571" class="Keyword">private</a>
|
|
<a id="581" class="Keyword">variable</a>
|
|
<a id="594" href="Algebra.Morphism.html#594" class="Generalizable">a</a> <a id="596" href="Algebra.Morphism.html#596" class="Generalizable">b</a> <a id="598" href="Algebra.Morphism.html#598" class="Generalizable">ℓ₁</a> <a id="601" href="Algebra.Morphism.html#601" class="Generalizable">ℓ₂</a> <a id="604" class="Symbol">:</a> <a id="606" href="Agda.Primitive.html#742" class="Postulate">Level</a>
|
|
<a id="616" href="Algebra.Morphism.html#616" class="Generalizable">A</a> <a id="618" class="Symbol">:</a> <a id="620" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="624" href="Algebra.Morphism.html#594" class="Generalizable">a</a>
|
|
<a id="630" href="Algebra.Morphism.html#630" class="Generalizable">B</a> <a id="632" class="Symbol">:</a> <a id="634" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="638" href="Algebra.Morphism.html#596" class="Generalizable">b</a>
|
|
|
|
<a id="641" class="Comment">------------------------------------------------------------------------</a>
|
|
<a id="714" class="Comment">-- Re-export</a>
|
|
|
|
<a id="728" class="Keyword">module</a> <a id="Definitions"></a><a id="735" href="Algebra.Morphism.html#735" class="Module">Definitions</a> <a id="747" class="Symbol">{</a><a id="748" href="Algebra.Morphism.html#748" class="Bound">a</a> <a id="750" href="Algebra.Morphism.html#750" class="Bound">b</a> <a id="752" href="Algebra.Morphism.html#752" class="Bound">ℓ₁</a><a id="754" class="Symbol">}</a> <a id="756" class="Symbol">(</a><a id="757" href="Algebra.Morphism.html#757" class="Bound">A</a> <a id="759" class="Symbol">:</a> <a id="761" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="765" href="Algebra.Morphism.html#748" class="Bound">a</a><a id="766" class="Symbol">)</a> <a id="768" class="Symbol">(</a><a id="769" href="Algebra.Morphism.html#769" class="Bound">B</a> <a id="771" class="Symbol">:</a> <a id="773" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="777" href="Algebra.Morphism.html#750" class="Bound">b</a><a id="778" class="Symbol">)</a> <a id="780" class="Symbol">(</a><a id="781" href="Algebra.Morphism.html#781" class="Bound Operator">_≈_</a> <a id="785" class="Symbol">:</a> <a id="787" href="Relation.Binary.Core.html#896" class="Function">Rel</a> <a id="791" href="Algebra.Morphism.html#769" class="Bound">B</a> <a id="793" href="Algebra.Morphism.html#752" class="Bound">ℓ₁</a><a id="795" class="Symbol">)</a> <a id="797" class="Keyword">where</a>
|
|
<a id="805" class="Keyword">open</a> <a id="810" href="Algebra.Morphism.Definitions.html" class="Module">MorphismDefinitions</a> <a id="830" href="Algebra.Morphism.html#757" class="Bound">A</a> <a id="832" href="Algebra.Morphism.html#769" class="Bound">B</a> <a id="834" href="Algebra.Morphism.html#781" class="Bound Operator">_≈_</a> <a id="838" class="Keyword">public</a>
|
|
|
|
<a id="846" class="Keyword">open</a> <a id="851" class="Keyword">import</a> <a id="858" href="Algebra.Morphism.Structures.html" class="Module">Algebra.Morphism.Structures</a> <a id="886" class="Keyword">public</a>
|
|
|
|
|
|
<a id="895" class="Comment">------------------------------------------------------------------------</a>
|
|
<a id="968" class="Comment">-- DEPRECATED</a>
|
|
<a id="982" class="Comment">------------------------------------------------------------------------</a>
|
|
<a id="1055" class="Comment">-- Please use the new definitions re-exported from</a>
|
|
<a id="1106" class="Comment">-- `Algebra.Morphism.Structures` as continuing support for the below is</a>
|
|
<a id="1178" class="Comment">-- no guaranteed.</a>
|
|
|
|
<a id="1197" class="Comment">-- Version 1.5</a>
|
|
|
|
<a id="1213" class="Keyword">module</a> <a id="1220" href="Algebra.Morphism.html#1220" class="Module">_</a> <a id="1222" class="Symbol">{</a><a id="1223" href="Algebra.Morphism.html#1223" class="Bound">c₁</a> <a id="1226" href="Algebra.Morphism.html#1226" class="Bound">ℓ₁</a> <a id="1229" href="Algebra.Morphism.html#1229" class="Bound">c₂</a> <a id="1232" href="Algebra.Morphism.html#1232" class="Bound">ℓ₂</a><a id="1234" class="Symbol">}</a>
|
|
<a id="1245" class="Symbol">(</a><a id="1246" href="Algebra.Morphism.html#1246" class="Bound">From</a> <a id="1251" class="Symbol">:</a> <a id="1253" href="Algebra.Bundles.html#4139" class="Record">Semigroup</a> <a id="1263" href="Algebra.Morphism.html#1223" class="Bound">c₁</a> <a id="1266" href="Algebra.Morphism.html#1226" class="Bound">ℓ₁</a><a id="1268" class="Symbol">)</a>
|
|
<a id="1279" class="Symbol">(</a><a id="1280" href="Algebra.Morphism.html#1280" class="Bound">To</a> <a id="1285" class="Symbol">:</a> <a id="1287" href="Algebra.Bundles.html#4139" class="Record">Semigroup</a> <a id="1297" href="Algebra.Morphism.html#1229" class="Bound">c₂</a> <a id="1300" href="Algebra.Morphism.html#1232" class="Bound">ℓ₂</a><a id="1302" class="Symbol">)</a> <a id="1304" class="Keyword">where</a>
|
|
|
|
<a id="1313" class="Keyword">private</a>
|
|
<a id="1325" class="Keyword">module</a> <a id="1332" href="Algebra.Morphism.html#1332" class="Module">F</a> <a id="1334" class="Symbol">=</a> <a id="1336" href="Algebra.Bundles.html#4139" class="Module">Semigroup</a> <a id="1346" href="Algebra.Morphism.html#1246" class="Bound">From</a>
|
|
<a id="1355" class="Keyword">module</a> <a id="1362" href="Algebra.Morphism.html#1362" class="Module">T</a> <a id="1364" class="Symbol">=</a> <a id="1366" href="Algebra.Bundles.html#4139" class="Module">Semigroup</a> <a id="1376" href="Algebra.Morphism.html#1280" class="Bound">To</a>
|
|
<a id="1381" class="Keyword">open</a> <a id="1386" href="Algebra.Morphism.html#735" class="Module">Definitions</a> <a id="1398" href="Algebra.Bundles.html#4221" class="Function">F.Carrier</a> <a id="1408" href="Algebra.Bundles.html#4221" class="Field">T.Carrier</a> <a id="1418" href="Algebra.Bundles.html#4245" class="Field Operator">T._≈_</a>
|
|
|
|
<a id="1427" class="Keyword">record</a> <a id="1434" href="Algebra.Morphism.html#1434" class="Record">IsSemigroupMorphism</a> <a id="1454" class="Symbol">(</a><a id="1455" href="Algebra.Morphism.html#1455" class="Bound Operator">⟦_⟧</a> <a id="1459" class="Symbol">:</a> <a id="1461" href="Algebra.Morphism.Definitions.html#1236" class="Function">Morphism</a><a id="1469" class="Symbol">)</a> <a id="1471" class="Symbol">:</a>
|
|
<a id="1482" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1486" class="Symbol">(</a><a id="1487" href="Algebra.Morphism.html#1223" class="Bound">c₁</a> <a id="1490" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1492" href="Algebra.Morphism.html#1226" class="Bound">ℓ₁</a> <a id="1495" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1497" href="Algebra.Morphism.html#1229" class="Bound">c₂</a> <a id="1500" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="1502" href="Algebra.Morphism.html#1232" class="Bound">ℓ₂</a><a id="1504" class="Symbol">)</a> <a id="1506" class="Keyword">where</a>
|
|
<a id="1516" class="Keyword">field</a>
|
|
<a id="1528" href="Algebra.Morphism.html#1528" class="Field">⟦⟧-cong</a> <a id="1536" class="Symbol">:</a> <a id="1538" href="Algebra.Morphism.html#1455" class="Bound Operator">⟦_⟧</a> <a id="1542" href="Relation.Binary.Core.html#1577" class="Function Operator">Preserves</a> <a id="1552" href="Algebra.Bundles.html#4245" class="Function Operator">F._≈_</a> <a id="1558" href="Relation.Binary.Core.html#1577" class="Function Operator">⟶</a> <a id="1560" href="Algebra.Bundles.html#4245" class="Field Operator">T._≈_</a>
|
|
<a id="1572" href="Algebra.Morphism.html#1572" class="Field">∙-homo</a> <a id="1580" class="Symbol">:</a> <a id="1582" href="Algebra.Morphism.Definitions.html#852" class="Function">Homomorphic₂</a> <a id="1595" href="Algebra.Morphism.html#1455" class="Bound Operator">⟦_⟧</a> <a id="1599" href="Algebra.Bundles.html#4277" class="Function Operator">F._∙_</a> <a id="1605" href="Algebra.Bundles.html#4277" class="Field Operator">T._∙_</a>
|
|
|
|
<a id="1614" href="Algebra.Morphism.html#1614" class="Function">IsSemigroupMorphism-syntax</a> <a id="1641" class="Symbol">=</a> <a id="1643" href="Algebra.Morphism.html#1434" class="Record">IsSemigroupMorphism</a>
|
|
<a id="1665" class="Keyword">syntax</a> <a id="1672" href="Algebra.Morphism.html#1614" class="Function">IsSemigroupMorphism-syntax</a> <a id="1699" class="Bound">From</a> <a id="1704" class="Bound">To</a> <a id="1707" class="Bound">F</a> <a id="1709" class="Symbol">=</a> <a id="1711" class="Bound">F</a> <a id="1713" class="Function">Is</a> <a id="1716" class="Bound">From</a> <a id="1721" class="Function">-Semigroup⟶</a> <a id="1733" class="Bound">To</a>
|
|
|
|
<a id="1737" class="Keyword">module</a> <a id="1744" href="Algebra.Morphism.html#1744" class="Module">_</a> <a id="1746" class="Symbol">{</a><a id="1747" href="Algebra.Morphism.html#1747" class="Bound">c₁</a> <a id="1750" href="Algebra.Morphism.html#1750" class="Bound">ℓ₁</a> <a id="1753" href="Algebra.Morphism.html#1753" class="Bound">c₂</a> <a id="1756" href="Algebra.Morphism.html#1756" class="Bound">ℓ₂</a><a id="1758" class="Symbol">}</a>
|
|
<a id="1769" class="Symbol">(</a><a id="1770" href="Algebra.Morphism.html#1770" class="Bound">From</a> <a id="1775" class="Symbol">:</a> <a id="1777" href="Algebra.Bundles.html#6051" class="Record">Monoid</a> <a id="1784" href="Algebra.Morphism.html#1747" class="Bound">c₁</a> <a id="1787" href="Algebra.Morphism.html#1750" class="Bound">ℓ₁</a><a id="1789" class="Symbol">)</a>
|
|
<a id="1800" class="Symbol">(</a><a id="1801" href="Algebra.Morphism.html#1801" class="Bound">To</a> <a id="1806" class="Symbol">:</a> <a id="1808" href="Algebra.Bundles.html#6051" class="Record">Monoid</a> <a id="1815" href="Algebra.Morphism.html#1753" class="Bound">c₂</a> <a id="1818" href="Algebra.Morphism.html#1756" class="Bound">ℓ₂</a><a id="1820" class="Symbol">)</a> <a id="1822" class="Keyword">where</a>
|
|
|
|
<a id="1831" class="Keyword">private</a>
|
|
<a id="1843" class="Keyword">module</a> <a id="1850" href="Algebra.Morphism.html#1850" class="Module">F</a> <a id="1852" class="Symbol">=</a> <a id="1854" href="Algebra.Bundles.html#6051" class="Module">Monoid</a> <a id="1861" href="Algebra.Morphism.html#1770" class="Bound">From</a>
|
|
<a id="1870" class="Keyword">module</a> <a id="1877" href="Algebra.Morphism.html#1877" class="Module">T</a> <a id="1879" class="Symbol">=</a> <a id="1881" href="Algebra.Bundles.html#6051" class="Module">Monoid</a> <a id="1888" href="Algebra.Morphism.html#1801" class="Bound">To</a>
|
|
<a id="1893" class="Keyword">open</a> <a id="1898" href="Algebra.Morphism.html#735" class="Module">Definitions</a> <a id="1910" href="Algebra.Bundles.html#6130" class="Function">F.Carrier</a> <a id="1920" href="Algebra.Bundles.html#6130" class="Field">T.Carrier</a> <a id="1930" href="Algebra.Bundles.html#6151" class="Field Operator">T._≈_</a>
|
|
|
|
<a id="1939" class="Keyword">record</a> <a id="1946" href="Algebra.Morphism.html#1946" class="Record">IsMonoidMorphism</a> <a id="1963" class="Symbol">(</a><a id="1964" href="Algebra.Morphism.html#1964" class="Bound Operator">⟦_⟧</a> <a id="1968" class="Symbol">:</a> <a id="1970" href="Algebra.Morphism.Definitions.html#1236" class="Function">Morphism</a><a id="1978" class="Symbol">)</a> <a id="1980" class="Symbol">:</a>
|
|
<a id="1991" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="1995" class="Symbol">(</a><a id="1996" href="Algebra.Morphism.html#1747" class="Bound">c₁</a> <a id="1999" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2001" href="Algebra.Morphism.html#1750" class="Bound">ℓ₁</a> <a id="2004" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2006" href="Algebra.Morphism.html#1753" class="Bound">c₂</a> <a id="2009" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2011" href="Algebra.Morphism.html#1756" class="Bound">ℓ₂</a><a id="2013" class="Symbol">)</a> <a id="2015" class="Keyword">where</a>
|
|
<a id="2025" class="Keyword">field</a>
|
|
<a id="2037" href="Algebra.Morphism.html#2037" class="Field">sm-homo</a> <a id="2045" class="Symbol">:</a> <a id="2047" href="Algebra.Morphism.html#1434" class="Record">IsSemigroupMorphism</a> <a id="2067" href="Algebra.Bundles.html#6296" class="Function">F.semigroup</a> <a id="2079" href="Algebra.Bundles.html#6296" class="Function">T.semigroup</a> <a id="2091" href="Algebra.Morphism.html#1964" class="Bound Operator">⟦_⟧</a>
|
|
<a id="2101" href="Algebra.Morphism.html#2101" class="Field">ε-homo</a> <a id="2109" class="Symbol">:</a> <a id="2111" href="Algebra.Morphism.Definitions.html#680" class="Function">Homomorphic₀</a> <a id="2124" href="Algebra.Morphism.html#1964" class="Bound Operator">⟦_⟧</a> <a id="2128" href="Algebra.Bundles.html#6207" class="Function">F.ε</a> <a id="2132" href="Algebra.Bundles.html#6207" class="Field">T.ε</a>
|
|
|
|
<a id="2141" class="Keyword">open</a> <a id="2146" href="Algebra.Morphism.html#1434" class="Module">IsSemigroupMorphism</a> <a id="2166" href="Algebra.Morphism.html#2037" class="Field">sm-homo</a> <a id="2174" class="Keyword">public</a>
|
|
|
|
<a id="2184" href="Algebra.Morphism.html#2184" class="Function">IsMonoidMorphism-syntax</a> <a id="2208" class="Symbol">=</a> <a id="2210" href="Algebra.Morphism.html#1946" class="Record">IsMonoidMorphism</a>
|
|
<a id="2229" class="Keyword">syntax</a> <a id="2236" href="Algebra.Morphism.html#2184" class="Function">IsMonoidMorphism-syntax</a> <a id="2260" class="Bound">From</a> <a id="2265" class="Bound">To</a> <a id="2268" class="Bound">F</a> <a id="2270" class="Symbol">=</a> <a id="2272" class="Bound">F</a> <a id="2274" class="Function">Is</a> <a id="2277" class="Bound">From</a> <a id="2282" class="Function">-Monoid⟶</a> <a id="2291" class="Bound">To</a>
|
|
|
|
<a id="2295" class="Keyword">module</a> <a id="2302" href="Algebra.Morphism.html#2302" class="Module">_</a> <a id="2304" class="Symbol">{</a><a id="2305" href="Algebra.Morphism.html#2305" class="Bound">c₁</a> <a id="2308" href="Algebra.Morphism.html#2308" class="Bound">ℓ₁</a> <a id="2311" href="Algebra.Morphism.html#2311" class="Bound">c₂</a> <a id="2314" href="Algebra.Morphism.html#2314" class="Bound">ℓ₂</a><a id="2316" class="Symbol">}</a>
|
|
<a id="2327" class="Symbol">(</a><a id="2328" href="Algebra.Morphism.html#2328" class="Bound">From</a> <a id="2333" class="Symbol">:</a> <a id="2335" href="Algebra.Bundles.html#6622" class="Record">CommutativeMonoid</a> <a id="2353" href="Algebra.Morphism.html#2305" class="Bound">c₁</a> <a id="2356" href="Algebra.Morphism.html#2308" class="Bound">ℓ₁</a><a id="2358" class="Symbol">)</a>
|
|
<a id="2369" class="Symbol">(</a><a id="2370" href="Algebra.Morphism.html#2370" class="Bound">To</a> <a id="2375" class="Symbol">:</a> <a id="2377" href="Algebra.Bundles.html#6622" class="Record">CommutativeMonoid</a> <a id="2395" href="Algebra.Morphism.html#2311" class="Bound">c₂</a> <a id="2398" href="Algebra.Morphism.html#2314" class="Bound">ℓ₂</a><a id="2400" class="Symbol">)</a> <a id="2402" class="Keyword">where</a>
|
|
|
|
<a id="2411" class="Keyword">private</a>
|
|
<a id="2423" class="Keyword">module</a> <a id="2430" href="Algebra.Morphism.html#2430" class="Module">F</a> <a id="2432" class="Symbol">=</a> <a id="2434" href="Algebra.Bundles.html#6622" class="Module">CommutativeMonoid</a> <a id="2452" href="Algebra.Morphism.html#2328" class="Bound">From</a>
|
|
<a id="2461" class="Keyword">module</a> <a id="2468" href="Algebra.Morphism.html#2468" class="Module">T</a> <a id="2470" class="Symbol">=</a> <a id="2472" href="Algebra.Bundles.html#6622" class="Module">CommutativeMonoid</a> <a id="2490" href="Algebra.Morphism.html#2370" class="Bound">To</a>
|
|
<a id="2495" class="Keyword">open</a> <a id="2500" href="Algebra.Morphism.html#735" class="Module">Definitions</a> <a id="2512" href="Algebra.Bundles.html#6712" class="Function">F.Carrier</a> <a id="2522" href="Algebra.Bundles.html#6712" class="Field">T.Carrier</a> <a id="2532" href="Algebra.Bundles.html#6744" class="Field Operator">T._≈_</a>
|
|
|
|
<a id="2541" class="Keyword">record</a> <a id="2548" href="Algebra.Morphism.html#2548" class="Record">IsCommutativeMonoidMorphism</a> <a id="2576" class="Symbol">(</a><a id="2577" href="Algebra.Morphism.html#2577" class="Bound Operator">⟦_⟧</a> <a id="2581" class="Symbol">:</a> <a id="2583" href="Algebra.Morphism.Definitions.html#1236" class="Function">Morphism</a><a id="2591" class="Symbol">)</a> <a id="2593" class="Symbol">:</a>
|
|
<a id="2604" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="2608" class="Symbol">(</a><a id="2609" href="Algebra.Morphism.html#2305" class="Bound">c₁</a> <a id="2612" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2614" href="Algebra.Morphism.html#2308" class="Bound">ℓ₁</a> <a id="2617" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2619" href="Algebra.Morphism.html#2311" class="Bound">c₂</a> <a id="2622" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="2624" href="Algebra.Morphism.html#2314" class="Bound">ℓ₂</a><a id="2626" class="Symbol">)</a> <a id="2628" class="Keyword">where</a>
|
|
<a id="2638" class="Keyword">field</a>
|
|
<a id="2650" href="Algebra.Morphism.html#2650" class="Field">mn-homo</a> <a id="2658" class="Symbol">:</a> <a id="2660" href="Algebra.Morphism.html#1946" class="Record">IsMonoidMorphism</a> <a id="2677" href="Algebra.Bundles.html#6966" class="Function">F.monoid</a> <a id="2686" href="Algebra.Bundles.html#6966" class="Function">T.monoid</a> <a id="2695" href="Algebra.Morphism.html#2577" class="Bound Operator">⟦_⟧</a>
|
|
|
|
<a id="2704" class="Keyword">open</a> <a id="2709" href="Algebra.Morphism.html#1946" class="Module">IsMonoidMorphism</a> <a id="2726" href="Algebra.Morphism.html#2650" class="Field">mn-homo</a> <a id="2734" class="Keyword">public</a>
|
|
|
|
<a id="2744" href="Algebra.Morphism.html#2744" class="Function">IsCommutativeMonoidMorphism-syntax</a> <a id="2779" class="Symbol">=</a> <a id="2781" href="Algebra.Morphism.html#2548" class="Record">IsCommutativeMonoidMorphism</a>
|
|
<a id="2811" class="Keyword">syntax</a> <a id="2818" href="Algebra.Morphism.html#2744" class="Function">IsCommutativeMonoidMorphism-syntax</a> <a id="2853" class="Bound">From</a> <a id="2858" class="Bound">To</a> <a id="2861" class="Bound">F</a> <a id="2863" class="Symbol">=</a> <a id="2865" class="Bound">F</a> <a id="2867" class="Function">Is</a> <a id="2870" class="Bound">From</a> <a id="2875" class="Function">-CommutativeMonoid⟶</a> <a id="2895" class="Bound">To</a>
|
|
|
|
<a id="2899" class="Keyword">module</a> <a id="2906" href="Algebra.Morphism.html#2906" class="Module">_</a> <a id="2908" class="Symbol">{</a><a id="2909" href="Algebra.Morphism.html#2909" class="Bound">c₁</a> <a id="2912" href="Algebra.Morphism.html#2912" class="Bound">ℓ₁</a> <a id="2915" href="Algebra.Morphism.html#2915" class="Bound">c₂</a> <a id="2918" href="Algebra.Morphism.html#2918" class="Bound">ℓ₂</a><a id="2920" class="Symbol">}</a>
|
|
<a id="2931" class="Symbol">(</a><a id="2932" href="Algebra.Morphism.html#2932" class="Bound">From</a> <a id="2937" class="Symbol">:</a> <a id="2939" href="Algebra.Bundles.html#7355" class="Record">IdempotentCommutativeMonoid</a> <a id="2967" href="Algebra.Morphism.html#2909" class="Bound">c₁</a> <a id="2970" href="Algebra.Morphism.html#2912" class="Bound">ℓ₁</a><a id="2972" class="Symbol">)</a>
|
|
<a id="2983" class="Symbol">(</a><a id="2984" href="Algebra.Morphism.html#2984" class="Bound">To</a> <a id="2989" class="Symbol">:</a> <a id="2991" href="Algebra.Bundles.html#7355" class="Record">IdempotentCommutativeMonoid</a> <a id="3019" href="Algebra.Morphism.html#2915" class="Bound">c₂</a> <a id="3022" href="Algebra.Morphism.html#2918" class="Bound">ℓ₂</a><a id="3024" class="Symbol">)</a> <a id="3026" class="Keyword">where</a>
|
|
|
|
<a id="3035" class="Keyword">private</a>
|
|
<a id="3047" class="Keyword">module</a> <a id="3054" href="Algebra.Morphism.html#3054" class="Module">F</a> <a id="3056" class="Symbol">=</a> <a id="3058" href="Algebra.Bundles.html#7355" class="Module">IdempotentCommutativeMonoid</a> <a id="3086" href="Algebra.Morphism.html#2932" class="Bound">From</a>
|
|
<a id="3095" class="Keyword">module</a> <a id="3102" href="Algebra.Morphism.html#3102" class="Module">T</a> <a id="3104" class="Symbol">=</a> <a id="3106" href="Algebra.Bundles.html#7355" class="Module">IdempotentCommutativeMonoid</a> <a id="3134" href="Algebra.Morphism.html#2984" class="Bound">To</a>
|
|
<a id="3139" class="Keyword">open</a> <a id="3144" href="Algebra.Morphism.html#735" class="Module">Definitions</a> <a id="3156" href="Algebra.Bundles.html#7455" class="Function">F.Carrier</a> <a id="3166" href="Algebra.Bundles.html#7455" class="Field">T.Carrier</a> <a id="3176" href="Algebra.Bundles.html#7497" class="Field Operator">T._≈_</a>
|
|
|
|
<a id="3185" class="Keyword">record</a> <a id="3192" href="Algebra.Morphism.html#3192" class="Record">IsIdempotentCommutativeMonoidMorphism</a> <a id="3230" class="Symbol">(</a><a id="3231" href="Algebra.Morphism.html#3231" class="Bound Operator">⟦_⟧</a> <a id="3235" class="Symbol">:</a> <a id="3237" href="Algebra.Morphism.Definitions.html#1236" class="Function">Morphism</a><a id="3245" class="Symbol">)</a> <a id="3247" class="Symbol">:</a>
|
|
<a id="3258" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="3262" class="Symbol">(</a><a id="3263" href="Algebra.Morphism.html#2909" class="Bound">c₁</a> <a id="3266" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="3268" href="Algebra.Morphism.html#2912" class="Bound">ℓ₁</a> <a id="3271" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="3273" href="Algebra.Morphism.html#2915" class="Bound">c₂</a> <a id="3276" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="3278" href="Algebra.Morphism.html#2918" class="Bound">ℓ₂</a><a id="3280" class="Symbol">)</a> <a id="3282" class="Keyword">where</a>
|
|
<a id="3292" class="Keyword">field</a>
|
|
<a id="3304" href="Algebra.Morphism.html#3304" class="Field">mn-homo</a> <a id="3312" class="Symbol">:</a> <a id="3314" href="Algebra.Morphism.html#1946" class="Record">IsMonoidMorphism</a> <a id="3331" href="Algebra.Bundles.html#6966" class="Function">F.monoid</a> <a id="3340" href="Algebra.Bundles.html#6966" class="Function">T.monoid</a> <a id="3349" href="Algebra.Morphism.html#3231" class="Bound Operator">⟦_⟧</a>
|
|
|
|
<a id="3358" class="Keyword">open</a> <a id="3363" href="Algebra.Morphism.html#1946" class="Module">IsMonoidMorphism</a> <a id="3380" href="Algebra.Morphism.html#3304" class="Field">mn-homo</a> <a id="3388" class="Keyword">public</a>
|
|
|
|
<a id="3400" href="Algebra.Morphism.html#3400" class="Function">isCommutativeMonoidMorphism</a> <a id="3428" class="Symbol">:</a>
|
|
<a id="3436" href="Algebra.Morphism.html#2548" class="Record">IsCommutativeMonoidMorphism</a> <a id="3464" href="Algebra.Bundles.html#7789" class="Function">F.commutativeMonoid</a> <a id="3484" href="Algebra.Bundles.html#7789" class="Function">T.commutativeMonoid</a> <a id="3504" href="Algebra.Morphism.html#3231" class="Bound Operator">⟦_⟧</a>
|
|
<a id="3512" href="Algebra.Morphism.html#3400" class="Function">isCommutativeMonoidMorphism</a> <a id="3540" class="Symbol">=</a> <a id="3542" class="Keyword">record</a> <a id="3549" class="Symbol">{</a> <a id="3551" href="Algebra.Morphism.html#2650" class="Field">mn-homo</a> <a id="3559" class="Symbol">=</a> <a id="3561" href="Algebra.Morphism.html#3304" class="Field">mn-homo</a> <a id="3569" class="Symbol">}</a>
|
|
|
|
<a id="3574" href="Algebra.Morphism.html#3574" class="Function">IsIdempotentCommutativeMonoidMorphism-syntax</a> <a id="3619" class="Symbol">=</a> <a id="3621" href="Algebra.Morphism.html#3192" class="Record">IsIdempotentCommutativeMonoidMorphism</a>
|
|
<a id="3661" class="Keyword">syntax</a> <a id="3668" href="Algebra.Morphism.html#3574" class="Function">IsIdempotentCommutativeMonoidMorphism-syntax</a> <a id="3713" class="Bound">From</a> <a id="3718" class="Bound">To</a> <a id="3721" class="Bound">F</a> <a id="3723" class="Symbol">=</a> <a id="3725" class="Bound">F</a> <a id="3727" class="Function">Is</a> <a id="3730" class="Bound">From</a> <a id="3735" class="Function">-IdempotentCommutativeMonoid⟶</a> <a id="3765" class="Bound">To</a>
|
|
|
|
<a id="3769" class="Keyword">module</a> <a id="3776" href="Algebra.Morphism.html#3776" class="Module">_</a> <a id="3778" class="Symbol">{</a><a id="3779" href="Algebra.Morphism.html#3779" class="Bound">c₁</a> <a id="3782" href="Algebra.Morphism.html#3782" class="Bound">ℓ₁</a> <a id="3785" href="Algebra.Morphism.html#3785" class="Bound">c₂</a> <a id="3788" href="Algebra.Morphism.html#3788" class="Bound">ℓ₂</a><a id="3790" class="Symbol">}</a>
|
|
<a id="3801" class="Symbol">(</a><a id="3802" href="Algebra.Morphism.html#3802" class="Bound">From</a> <a id="3807" class="Symbol">:</a> <a id="3809" href="Algebra.Bundles.html#9767" class="Record">Group</a> <a id="3815" href="Algebra.Morphism.html#3779" class="Bound">c₁</a> <a id="3818" href="Algebra.Morphism.html#3782" class="Bound">ℓ₁</a><a id="3820" class="Symbol">)</a>
|
|
<a id="3831" class="Symbol">(</a><a id="3832" href="Algebra.Morphism.html#3832" class="Bound">To</a> <a id="3837" class="Symbol">:</a> <a id="3839" href="Algebra.Bundles.html#9767" class="Record">Group</a> <a id="3845" href="Algebra.Morphism.html#3785" class="Bound">c₂</a> <a id="3848" href="Algebra.Morphism.html#3788" class="Bound">ℓ₂</a><a id="3850" class="Symbol">)</a> <a id="3852" class="Keyword">where</a>
|
|
|
|
<a id="3861" class="Keyword">private</a>
|
|
<a id="3873" class="Keyword">module</a> <a id="3880" href="Algebra.Morphism.html#3880" class="Module">F</a> <a id="3882" class="Symbol">=</a> <a id="3884" href="Algebra.Bundles.html#9767" class="Module">Group</a> <a id="3890" href="Algebra.Morphism.html#3802" class="Bound">From</a>
|
|
<a id="3899" class="Keyword">module</a> <a id="3906" href="Algebra.Morphism.html#3906" class="Module">T</a> <a id="3908" class="Symbol">=</a> <a id="3910" href="Algebra.Bundles.html#9767" class="Module">Group</a> <a id="3916" href="Algebra.Morphism.html#3832" class="Bound">To</a>
|
|
<a id="3921" class="Keyword">open</a> <a id="3926" href="Algebra.Morphism.html#735" class="Module">Definitions</a> <a id="3938" href="Algebra.Bundles.html#9860" class="Function">F.Carrier</a> <a id="3948" href="Algebra.Bundles.html#9860" class="Field">T.Carrier</a> <a id="3958" href="Algebra.Bundles.html#9880" class="Field Operator">T._≈_</a>
|
|
|
|
<a id="3967" class="Keyword">record</a> <a id="3974" href="Algebra.Morphism.html#3974" class="Record">IsGroupMorphism</a> <a id="3990" class="Symbol">(</a><a id="3991" href="Algebra.Morphism.html#3991" class="Bound Operator">⟦_⟧</a> <a id="3995" class="Symbol">:</a> <a id="3997" href="Algebra.Morphism.Definitions.html#1236" class="Function">Morphism</a><a id="4005" class="Symbol">)</a> <a id="4007" class="Symbol">:</a>
|
|
<a id="4018" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="4022" class="Symbol">(</a><a id="4023" href="Algebra.Morphism.html#3779" class="Bound">c₁</a> <a id="4026" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="4028" href="Algebra.Morphism.html#3782" class="Bound">ℓ₁</a> <a id="4031" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="4033" href="Algebra.Morphism.html#3785" class="Bound">c₂</a> <a id="4036" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="4038" href="Algebra.Morphism.html#3788" class="Bound">ℓ₂</a><a id="4040" class="Symbol">)</a> <a id="4042" class="Keyword">where</a>
|
|
<a id="4052" class="Keyword">field</a>
|
|
<a id="4064" href="Algebra.Morphism.html#4064" class="Field">mn-homo</a> <a id="4072" class="Symbol">:</a> <a id="4074" href="Algebra.Morphism.html#1946" class="Record">IsMonoidMorphism</a> <a id="4091" href="Algebra.Bundles.html#10137" class="Function">F.monoid</a> <a id="4100" href="Algebra.Bundles.html#10137" class="Function">T.monoid</a> <a id="4109" href="Algebra.Morphism.html#3991" class="Bound Operator">⟦_⟧</a>
|
|
|
|
<a id="4118" class="Keyword">open</a> <a id="4123" href="Algebra.Morphism.html#1946" class="Module">IsMonoidMorphism</a> <a id="4140" href="Algebra.Morphism.html#4064" class="Field">mn-homo</a> <a id="4148" class="Keyword">public</a>
|
|
|
|
<a id="4160" href="Algebra.Morphism.html#4160" class="Function">⁻¹-homo</a> <a id="4168" class="Symbol">:</a> <a id="4170" href="Algebra.Morphism.Definitions.html#753" class="Function">Homomorphic₁</a> <a id="4183" href="Algebra.Morphism.html#3991" class="Bound Operator">⟦_⟧</a> <a id="4187" href="Algebra.Bundles.html#9956" class="Function Operator">F._⁻¹</a> <a id="4193" href="Algebra.Bundles.html#9956" class="Field Operator">T._⁻¹</a>
|
|
<a id="4203" href="Algebra.Morphism.html#4160" class="Function">⁻¹-homo</a> <a id="4211" href="Algebra.Morphism.html#4211" class="Bound">x</a> <a id="4213" class="Symbol">=</a> <a id="4215" class="Keyword">let</a> <a id="4219" class="Keyword">open</a> <a id="4224" href="Relation.Binary.Reasoning.Setoid.html" class="Module">EqR</a> <a id="4228" href="Algebra.Structures.html#1390" class="Function">T.setoid</a> <a id="4237" class="Keyword">in</a> <a id="4240" href="Algebra.Structures.html#6570" class="Function">T.uniqueˡ-⁻¹</a> <a id="4253" href="Algebra.Morphism.html#3991" class="Bound Operator">⟦</a> <a id="4255" href="Algebra.Morphism.html#4211" class="Bound">x</a> <a id="4257" href="Algebra.Bundles.html#9956" class="Function Operator">F.⁻¹</a> <a id="4262" href="Algebra.Morphism.html#3991" class="Bound Operator">⟧</a> <a id="4264" href="Algebra.Morphism.html#3991" class="Bound Operator">⟦</a> <a id="4266" href="Algebra.Morphism.html#4211" class="Bound">x</a> <a id="4268" href="Algebra.Morphism.html#3991" class="Bound Operator">⟧</a> <a id="4270" href="Function.Base.html#1974" class="Function Operator">$</a> <a id="4272" href="Relation.Binary.Reasoning.Syntax.html#1510" class="Function Operator">begin</a>
|
|
<a id="4284" href="Algebra.Morphism.html#3991" class="Bound Operator">⟦</a> <a id="4286" href="Algebra.Morphism.html#4211" class="Bound">x</a> <a id="4288" href="Algebra.Bundles.html#9956" class="Function Operator">F.⁻¹</a> <a id="4293" href="Algebra.Morphism.html#3991" class="Bound Operator">⟧</a> <a id="4295" href="Algebra.Bundles.html#9908" class="Field Operator">T.∙</a> <a id="4299" href="Algebra.Morphism.html#3991" class="Bound Operator">⟦</a> <a id="4301" href="Algebra.Morphism.html#4211" class="Bound">x</a> <a id="4303" href="Algebra.Morphism.html#3991" class="Bound Operator">⟧</a> <a id="4305" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4308" href="Relation.Binary.Structures.html#1622" class="Function">T.sym</a> <a id="4314" class="Symbol">(</a><a id="4315" href="Algebra.Morphism.html#1572" class="Function">∙-homo</a> <a id="4322" class="Symbol">(</a><a id="4323" href="Algebra.Morphism.html#4211" class="Bound">x</a> <a id="4325" href="Algebra.Bundles.html#9956" class="Function Operator">F.⁻¹</a><a id="4329" class="Symbol">)</a> <a id="4331" href="Algebra.Morphism.html#4211" class="Bound">x</a><a id="4332" class="Symbol">)</a> <a id="4334" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
|
<a id="4342" href="Algebra.Morphism.html#3991" class="Bound Operator">⟦</a> <a id="4344" href="Algebra.Morphism.html#4211" class="Bound">x</a> <a id="4346" href="Algebra.Bundles.html#9956" class="Function Operator">F.⁻¹</a> <a id="4351" href="Algebra.Bundles.html#9908" class="Function Operator">F.∙</a> <a id="4355" href="Algebra.Morphism.html#4211" class="Bound">x</a> <a id="4357" href="Algebra.Morphism.html#3991" class="Bound Operator">⟧</a> <a id="4363" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4366" href="Algebra.Morphism.html#1528" class="Function">⟦⟧-cong</a> <a id="4374" class="Symbol">(</a><a id="4375" href="Algebra.Structures.html#6443" class="Function">F.inverseˡ</a> <a id="4386" href="Algebra.Morphism.html#4211" class="Bound">x</a><a id="4387" class="Symbol">)</a> <a id="4389" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
|
<a id="4397" href="Algebra.Morphism.html#3991" class="Bound Operator">⟦</a> <a id="4399" href="Algebra.Bundles.html#9934" class="Function">F.ε</a> <a id="4403" href="Algebra.Morphism.html#3991" class="Bound Operator">⟧</a> <a id="4418" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">≈⟨</a> <a id="4421" href="Algebra.Morphism.html#2101" class="Function">ε-homo</a> <a id="4428" href="Relation.Binary.Reasoning.Syntax.html#7049" class="Function">⟩</a>
|
|
<a id="4436" href="Algebra.Bundles.html#9934" class="Field">T.ε</a> <a id="4440" href="Relation.Binary.Reasoning.Syntax.html#12283" class="Function Operator">∎</a>
|
|
|
|
<a id="4445" href="Algebra.Morphism.html#4445" class="Function">IsGroupMorphism-syntax</a> <a id="4468" class="Symbol">=</a> <a id="4470" href="Algebra.Morphism.html#3974" class="Record">IsGroupMorphism</a>
|
|
<a id="4488" class="Keyword">syntax</a> <a id="4495" href="Algebra.Morphism.html#4445" class="Function">IsGroupMorphism-syntax</a> <a id="4518" class="Bound">From</a> <a id="4523" class="Bound">To</a> <a id="4526" class="Bound">F</a> <a id="4528" class="Symbol">=</a> <a id="4530" class="Bound">F</a> <a id="4532" class="Function">Is</a> <a id="4535" class="Bound">From</a> <a id="4540" class="Function">-Group⟶</a> <a id="4548" class="Bound">To</a>
|
|
|
|
<a id="4552" class="Keyword">module</a> <a id="4559" href="Algebra.Morphism.html#4559" class="Module">_</a> <a id="4561" class="Symbol">{</a><a id="4562" href="Algebra.Morphism.html#4562" class="Bound">c₁</a> <a id="4565" href="Algebra.Morphism.html#4565" class="Bound">ℓ₁</a> <a id="4568" href="Algebra.Morphism.html#4568" class="Bound">c₂</a> <a id="4571" href="Algebra.Morphism.html#4571" class="Bound">ℓ₂</a><a id="4573" class="Symbol">}</a>
|
|
<a id="4584" class="Symbol">(</a><a id="4585" href="Algebra.Morphism.html#4585" class="Bound">From</a> <a id="4590" class="Symbol">:</a> <a id="4592" href="Algebra.Bundles.html#10570" class="Record">AbelianGroup</a> <a id="4605" href="Algebra.Morphism.html#4562" class="Bound">c₁</a> <a id="4608" href="Algebra.Morphism.html#4565" class="Bound">ℓ₁</a><a id="4610" class="Symbol">)</a>
|
|
<a id="4621" class="Symbol">(</a><a id="4622" href="Algebra.Morphism.html#4622" class="Bound">To</a> <a id="4627" class="Symbol">:</a> <a id="4629" href="Algebra.Bundles.html#10570" class="Record">AbelianGroup</a> <a id="4642" href="Algebra.Morphism.html#4568" class="Bound">c₂</a> <a id="4645" href="Algebra.Morphism.html#4571" class="Bound">ℓ₂</a><a id="4647" class="Symbol">)</a> <a id="4649" class="Keyword">where</a>
|
|
|
|
<a id="4658" class="Keyword">private</a>
|
|
<a id="4670" class="Keyword">module</a> <a id="4677" href="Algebra.Morphism.html#4677" class="Module">F</a> <a id="4679" class="Symbol">=</a> <a id="4681" href="Algebra.Bundles.html#10570" class="Module">AbelianGroup</a> <a id="4694" href="Algebra.Morphism.html#4585" class="Bound">From</a>
|
|
<a id="4703" class="Keyword">module</a> <a id="4710" href="Algebra.Morphism.html#4710" class="Module">T</a> <a id="4712" class="Symbol">=</a> <a id="4714" href="Algebra.Bundles.html#10570" class="Module">AbelianGroup</a> <a id="4727" href="Algebra.Morphism.html#4622" class="Bound">To</a>
|
|
<a id="4732" class="Keyword">open</a> <a id="4737" href="Algebra.Morphism.html#735" class="Module">Definitions</a> <a id="4749" href="Algebra.Bundles.html#10670" class="Function">F.Carrier</a> <a id="4759" href="Algebra.Bundles.html#10670" class="Field">T.Carrier</a> <a id="4769" href="Algebra.Bundles.html#10697" class="Field Operator">T._≈_</a>
|
|
|
|
<a id="4778" class="Keyword">record</a> <a id="4785" href="Algebra.Morphism.html#4785" class="Record">IsAbelianGroupMorphism</a> <a id="4808" class="Symbol">(</a><a id="4809" href="Algebra.Morphism.html#4809" class="Bound Operator">⟦_⟧</a> <a id="4813" class="Symbol">:</a> <a id="4815" href="Algebra.Morphism.Definitions.html#1236" class="Function">Morphism</a><a id="4823" class="Symbol">)</a> <a id="4825" class="Symbol">:</a>
|
|
<a id="4836" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="4840" class="Symbol">(</a><a id="4841" href="Algebra.Morphism.html#4562" class="Bound">c₁</a> <a id="4844" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="4846" href="Algebra.Morphism.html#4565" class="Bound">ℓ₁</a> <a id="4849" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="4851" href="Algebra.Morphism.html#4568" class="Bound">c₂</a> <a id="4854" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="4856" href="Algebra.Morphism.html#4571" class="Bound">ℓ₂</a><a id="4858" class="Symbol">)</a> <a id="4860" class="Keyword">where</a>
|
|
<a id="4870" class="Keyword">field</a>
|
|
<a id="4882" href="Algebra.Morphism.html#4882" class="Field">gp-homo</a> <a id="4890" class="Symbol">:</a> <a id="4892" href="Algebra.Morphism.html#3974" class="Record">IsGroupMorphism</a> <a id="4908" href="Algebra.Bundles.html#10921" class="Function">F.group</a> <a id="4916" href="Algebra.Bundles.html#10921" class="Function">T.group</a> <a id="4924" href="Algebra.Morphism.html#4809" class="Bound Operator">⟦_⟧</a>
|
|
|
|
<a id="4933" class="Keyword">open</a> <a id="4938" href="Algebra.Morphism.html#3974" class="Module">IsGroupMorphism</a> <a id="4954" href="Algebra.Morphism.html#4882" class="Field">gp-homo</a> <a id="4962" class="Keyword">public</a>
|
|
|
|
<a id="4972" href="Algebra.Morphism.html#4972" class="Function">IsAbelianGroupMorphism-syntax</a> <a id="5002" class="Symbol">=</a> <a id="5004" href="Algebra.Morphism.html#4785" class="Record">IsAbelianGroupMorphism</a>
|
|
<a id="5029" class="Keyword">syntax</a> <a id="5036" href="Algebra.Morphism.html#4972" class="Function">IsAbelianGroupMorphism-syntax</a> <a id="5066" class="Bound">From</a> <a id="5071" class="Bound">To</a> <a id="5074" class="Bound">F</a> <a id="5076" class="Symbol">=</a> <a id="5078" class="Bound">F</a> <a id="5080" class="Function">Is</a> <a id="5083" class="Bound">From</a> <a id="5088" class="Function">-AbelianGroup⟶</a> <a id="5103" class="Bound">To</a>
|
|
|
|
<a id="5107" class="Keyword">module</a> <a id="5114" href="Algebra.Morphism.html#5114" class="Module">_</a> <a id="5116" class="Symbol">{</a><a id="5117" href="Algebra.Morphism.html#5117" class="Bound">c₁</a> <a id="5120" href="Algebra.Morphism.html#5120" class="Bound">ℓ₁</a> <a id="5123" href="Algebra.Morphism.html#5123" class="Bound">c₂</a> <a id="5126" href="Algebra.Morphism.html#5126" class="Bound">ℓ₂</a><a id="5128" class="Symbol">}</a>
|
|
<a id="5139" class="Symbol">(</a><a id="5140" href="Algebra.Morphism.html#5140" class="Bound">From</a> <a id="5145" class="Symbol">:</a> <a id="5147" href="Algebra.Bundles.html#25949" class="Record">Ring</a> <a id="5152" href="Algebra.Morphism.html#5117" class="Bound">c₁</a> <a id="5155" href="Algebra.Morphism.html#5120" class="Bound">ℓ₁</a><a id="5157" class="Symbol">)</a>
|
|
<a id="5168" class="Symbol">(</a><a id="5169" href="Algebra.Morphism.html#5169" class="Bound">To</a> <a id="5174" class="Symbol">:</a> <a id="5176" href="Algebra.Bundles.html#25949" class="Record">Ring</a> <a id="5181" href="Algebra.Morphism.html#5123" class="Bound">c₂</a> <a id="5184" href="Algebra.Morphism.html#5126" class="Bound">ℓ₂</a><a id="5186" class="Symbol">)</a> <a id="5188" class="Keyword">where</a>
|
|
|
|
<a id="5197" class="Keyword">private</a>
|
|
<a id="5209" class="Keyword">module</a> <a id="5216" href="Algebra.Morphism.html#5216" class="Module">F</a> <a id="5218" class="Symbol">=</a> <a id="5220" href="Algebra.Bundles.html#25949" class="Module">Ring</a> <a id="5225" href="Algebra.Morphism.html#5140" class="Bound">From</a>
|
|
<a id="5234" class="Keyword">module</a> <a id="5241" href="Algebra.Morphism.html#5241" class="Module">T</a> <a id="5243" class="Symbol">=</a> <a id="5245" href="Algebra.Bundles.html#25949" class="Module">Ring</a> <a id="5250" href="Algebra.Morphism.html#5169" class="Bound">To</a>
|
|
<a id="5255" class="Keyword">open</a> <a id="5260" href="Algebra.Morphism.html#735" class="Module">Definitions</a> <a id="5272" href="Algebra.Bundles.html#26055" class="Function">F.Carrier</a> <a id="5282" href="Algebra.Bundles.html#26055" class="Field">T.Carrier</a> <a id="5292" href="Algebra.Bundles.html#26075" class="Field Operator">T._≈_</a>
|
|
|
|
<a id="5301" class="Keyword">record</a> <a id="5308" href="Algebra.Morphism.html#5308" class="Record">IsRingMorphism</a> <a id="5323" class="Symbol">(</a><a id="5324" href="Algebra.Morphism.html#5324" class="Bound Operator">⟦_⟧</a> <a id="5328" class="Symbol">:</a> <a id="5330" href="Algebra.Morphism.Definitions.html#1236" class="Function">Morphism</a><a id="5338" class="Symbol">)</a> <a id="5340" class="Symbol">:</a>
|
|
<a id="5351" href="Agda.Primitive.html#388" class="Primitive">Set</a> <a id="5355" class="Symbol">(</a><a id="5356" href="Algebra.Morphism.html#5117" class="Bound">c₁</a> <a id="5359" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="5361" href="Algebra.Morphism.html#5120" class="Bound">ℓ₁</a> <a id="5364" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="5366" href="Algebra.Morphism.html#5123" class="Bound">c₂</a> <a id="5369" href="Agda.Primitive.html#961" class="Primitive Operator">⊔</a> <a id="5371" href="Algebra.Morphism.html#5126" class="Bound">ℓ₂</a><a id="5373" class="Symbol">)</a> <a id="5375" class="Keyword">where</a>
|
|
<a id="5385" class="Keyword">field</a>
|
|
<a id="5397" href="Algebra.Morphism.html#5397" class="Field">+-abgp-homo</a> <a id="5409" class="Symbol">:</a> <a id="5411" href="Algebra.Morphism.html#5324" class="Bound Operator">⟦_⟧</a> <a id="5415" href="Algebra.Morphism.html#4972" class="Function">Is</a> <a id="5418" href="Algebra.Bundles.html#26295" class="Function">F.+-abelianGroup</a> <a id="5435" href="Algebra.Morphism.html#4972" class="Function">-AbelianGroup⟶</a> <a id="5450" href="Algebra.Bundles.html#26295" class="Function">T.+-abelianGroup</a>
|
|
<a id="5473" href="Algebra.Morphism.html#5473" class="Field">*-mn-homo</a> <a id="5485" class="Symbol">:</a> <a id="5487" href="Algebra.Morphism.html#5324" class="Bound Operator">⟦_⟧</a> <a id="5491" href="Algebra.Morphism.html#2184" class="Function">Is</a> <a id="5494" href="Algebra.Bundles.html#16091" class="Function">F.*-monoid</a> <a id="5505" href="Algebra.Morphism.html#2184" class="Function">-Monoid⟶</a> <a id="5514" href="Algebra.Bundles.html#16091" class="Function">T.*-monoid</a>
|
|
|
|
<a id="5528" href="Algebra.Morphism.html#5528" class="Function">IsRingMorphism-syntax</a> <a id="5550" class="Symbol">=</a> <a id="5552" href="Algebra.Morphism.html#5308" class="Record">IsRingMorphism</a>
|
|
<a id="5569" class="Keyword">syntax</a> <a id="5576" href="Algebra.Morphism.html#5528" class="Function">IsRingMorphism-syntax</a> <a id="5598" class="Bound">From</a> <a id="5603" class="Bound">To</a> <a id="5606" class="Bound">F</a> <a id="5608" class="Symbol">=</a> <a id="5610" class="Bound">F</a> <a id="5612" class="Function">Is</a> <a id="5615" class="Bound">From</a> <a id="5620" class="Function">-Ring⟶</a> <a id="5627" class="Bound">To</a>
|
|
|
|
<a id="5631" class="Symbol">{-#</a> <a id="5635" class="Keyword">WARNING_ON_USAGE</a> <a id="5652" class="Pragma">IsSemigroupMorphism</a>
|
|
<a id="5672" class="String">"Warning: IsSemigroupMorphism was deprecated in v1.5.
|
|
Please use IsSemigroupHomomorphism instead."</a>
|
|
<a id="5771" class="Symbol">#-}</a>
|
|
<a id="5775" class="Symbol">{-#</a> <a id="5779" class="Keyword">WARNING_ON_USAGE</a> <a id="5796" class="Pragma">IsMonoidMorphism</a>
|
|
<a id="5813" class="String">"Warning: IsMonoidMorphism was deprecated in v1.5.
|
|
Please use IsMonoidHomomorphism instead."</a>
|
|
<a id="5906" class="Symbol">#-}</a>
|
|
<a id="5910" class="Symbol">{-#</a> <a id="5914" class="Keyword">WARNING_ON_USAGE</a> <a id="5931" class="Pragma">IsCommutativeMonoidMorphism</a>
|
|
<a id="5959" class="String">"Warning: IsCommutativeMonoidMorphism was deprecated in v1.5.
|
|
Please use IsMonoidHomomorphism instead."</a>
|
|
<a id="6063" class="Symbol">#-}</a>
|
|
<a id="6067" class="Symbol">{-#</a> <a id="6071" class="Keyword">WARNING_ON_USAGE</a> <a id="6088" class="Pragma">IsIdempotentCommutativeMonoidMorphism</a>
|
|
<a id="6126" class="String">"Warning: IsIdempotentCommutativeMonoidMorphism was deprecated in v1.5.
|
|
Please use IsMonoidHomomorphism instead."</a>
|
|
<a id="6240" class="Symbol">#-}</a>
|
|
<a id="6244" class="Symbol">{-#</a> <a id="6248" class="Keyword">WARNING_ON_USAGE</a> <a id="6265" class="Pragma">IsGroupMorphism</a>
|
|
<a id="6281" class="String">"Warning: IsGroupMorphism was deprecated in v1.5.
|
|
Please use IsGroupHomomorphism instead."</a>
|
|
<a id="6372" class="Symbol">#-}</a>
|
|
<a id="6376" class="Symbol">{-#</a> <a id="6380" class="Keyword">WARNING_ON_USAGE</a> <a id="6397" class="Pragma">IsAbelianGroupMorphism</a>
|
|
<a id="6420" class="String">"Warning: IsAbelianGroupMorphism was deprecated in v1.5.
|
|
Please use IsGroupHomomorphism instead."</a>
|
|
<a id="6518" class="Symbol">#-}</a>
|
|
</pre></body></html> |