![]() |
Metamath
Proof Explorer Theorem List (p. 422 of 429) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Color key: | ![]() (1-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | ismgmd 42101* | Deduce a magma from its properties. (Contributed by AV, 25-Feb-2020.) |
⊢ (𝜑 → 𝐵 = (Base‘𝐺)) & ⊢ (𝜑 → 𝐺 ∈ 𝑉) & ⊢ (𝜑 → + = (+g‘𝐺)) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐵 ∧ 𝑦 ∈ 𝐵) → (𝑥 + 𝑦) ∈ 𝐵) ⇒ ⊢ (𝜑 → 𝐺 ∈ Mgm) | ||
Syntax | cmgmhm 42102 | Hom-set generator class for magmas. |
class MgmHom | ||
Syntax | csubmgm 42103 | Class function taking a magma to its lattice of submagmas. |
class SubMgm | ||
Definition | df-mgmhm 42104* | A magma homomorphism is a function on the base sets which preserves the binary operation. (Contributed by AV, 24-Feb-2020.) |
⊢ MgmHom = (𝑠 ∈ Mgm, 𝑡 ∈ Mgm ↦ {𝑓 ∈ ((Base‘𝑡) ↑𝑚 (Base‘𝑠)) ∣ ∀𝑥 ∈ (Base‘𝑠)∀𝑦 ∈ (Base‘𝑠)(𝑓‘(𝑥(+g‘𝑠)𝑦)) = ((𝑓‘𝑥)(+g‘𝑡)(𝑓‘𝑦))}) | ||
Definition | df-submgm 42105* | A submagma is a subset of a magma which is closed under the operation. Such subsets are themselves magmas. (Contributed by AV, 24-Feb-2020.) |
⊢ SubMgm = (𝑠 ∈ Mgm ↦ {𝑡 ∈ 𝒫 (Base‘𝑠) ∣ ∀𝑥 ∈ 𝑡 ∀𝑦 ∈ 𝑡 (𝑥(+g‘𝑠)𝑦) ∈ 𝑡}) | ||
Theorem | mgmhmrcl 42106 | Reverse closure of a magma homomorphism. (Contributed by AV, 24-Feb-2020.) |
⊢ (𝐹 ∈ (𝑆 MgmHom 𝑇) → (𝑆 ∈ Mgm ∧ 𝑇 ∈ Mgm)) | ||
Theorem | submgmrcl 42107 | Reverse closure for submagmas. (Contributed by AV, 24-Feb-2020.) |
⊢ (𝑆 ∈ (SubMgm‘𝑀) → 𝑀 ∈ Mgm) | ||
Theorem | ismgmhm 42108* | Property of a magma homomorphism. (Contributed by AV, 25-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑆) & ⊢ 𝐶 = (Base‘𝑇) & ⊢ + = (+g‘𝑆) & ⊢ ⨣ = (+g‘𝑇) ⇒ ⊢ (𝐹 ∈ (𝑆 MgmHom 𝑇) ↔ ((𝑆 ∈ Mgm ∧ 𝑇 ∈ Mgm) ∧ (𝐹:𝐵⟶𝐶 ∧ ∀𝑥 ∈ 𝐵 ∀𝑦 ∈ 𝐵 (𝐹‘(𝑥 + 𝑦)) = ((𝐹‘𝑥) ⨣ (𝐹‘𝑦))))) | ||
Theorem | mgmhmf 42109 | A magma homomorphism is a function. (Contributed by AV, 25-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑆) & ⊢ 𝐶 = (Base‘𝑇) ⇒ ⊢ (𝐹 ∈ (𝑆 MgmHom 𝑇) → 𝐹:𝐵⟶𝐶) | ||
Theorem | mgmhmpropd 42110* | Magma homomorphism depends only on the operation of structures. (Contributed by AV, 25-Feb-2020.) |
⊢ (𝜑 → 𝐵 = (Base‘𝐽)) & ⊢ (𝜑 → 𝐶 = (Base‘𝐾)) & ⊢ (𝜑 → 𝐵 = (Base‘𝐿)) & ⊢ (𝜑 → 𝐶 = (Base‘𝑀)) & ⊢ (𝜑 → 𝐵 ≠ ∅) & ⊢ (𝜑 → 𝐶 ≠ ∅) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝐵 ∧ 𝑦 ∈ 𝐵)) → (𝑥(+g‘𝐽)𝑦) = (𝑥(+g‘𝐿)𝑦)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝐶 ∧ 𝑦 ∈ 𝐶)) → (𝑥(+g‘𝐾)𝑦) = (𝑥(+g‘𝑀)𝑦)) ⇒ ⊢ (𝜑 → (𝐽 MgmHom 𝐾) = (𝐿 MgmHom 𝑀)) | ||
Theorem | mgmhmlin 42111 | A magma homomorphism preserves the binary operation. (Contributed by AV, 25-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑆) & ⊢ + = (+g‘𝑆) & ⊢ ⨣ = (+g‘𝑇) ⇒ ⊢ ((𝐹 ∈ (𝑆 MgmHom 𝑇) ∧ 𝑋 ∈ 𝐵 ∧ 𝑌 ∈ 𝐵) → (𝐹‘(𝑋 + 𝑌)) = ((𝐹‘𝑋) ⨣ (𝐹‘𝑌))) | ||
Theorem | mgmhmf1o 42112 | A magma homomorphism is bijective iff its converse is also a magma homomorphism. (Contributed by AV, 25-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑅) & ⊢ 𝐶 = (Base‘𝑆) ⇒ ⊢ (𝐹 ∈ (𝑅 MgmHom 𝑆) → (𝐹:𝐵–1-1-onto→𝐶 ↔ ◡𝐹 ∈ (𝑆 MgmHom 𝑅))) | ||
Theorem | idmgmhm 42113 | The identity homomorphism on a magma. (Contributed by AV, 27-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑀) ⇒ ⊢ (𝑀 ∈ Mgm → ( I ↾ 𝐵) ∈ (𝑀 MgmHom 𝑀)) | ||
Theorem | issubmgm 42114* | Expand definition of a submagma. (Contributed by AV, 25-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ + = (+g‘𝑀) ⇒ ⊢ (𝑀 ∈ Mgm → (𝑆 ∈ (SubMgm‘𝑀) ↔ (𝑆 ⊆ 𝐵 ∧ ∀𝑥 ∈ 𝑆 ∀𝑦 ∈ 𝑆 (𝑥 + 𝑦) ∈ 𝑆))) | ||
Theorem | issubmgm2 42115 | Submagmas are subsets that are also magmas. (Contributed by AV, 25-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ 𝐻 = (𝑀 ↾s 𝑆) ⇒ ⊢ (𝑀 ∈ Mgm → (𝑆 ∈ (SubMgm‘𝑀) ↔ (𝑆 ⊆ 𝐵 ∧ 𝐻 ∈ Mgm))) | ||
Theorem | rabsubmgmd 42116* | Deduction for proving that a restricted class abstraction is a submagma. (Contributed by AV, 26-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ + = (+g‘𝑀) & ⊢ (𝜑 → 𝑀 ∈ Mgm) & ⊢ ((𝜑 ∧ ((𝑥 ∈ 𝐵 ∧ 𝑦 ∈ 𝐵) ∧ (𝜃 ∧ 𝜏))) → 𝜂) & ⊢ (𝑧 = 𝑥 → (𝜓 ↔ 𝜃)) & ⊢ (𝑧 = 𝑦 → (𝜓 ↔ 𝜏)) & ⊢ (𝑧 = (𝑥 + 𝑦) → (𝜓 ↔ 𝜂)) ⇒ ⊢ (𝜑 → {𝑧 ∈ 𝐵 ∣ 𝜓} ∈ (SubMgm‘𝑀)) | ||
Theorem | submgmss 42117 | Submagmas are subsets of the base set. (Contributed by AV, 26-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑀) ⇒ ⊢ (𝑆 ∈ (SubMgm‘𝑀) → 𝑆 ⊆ 𝐵) | ||
Theorem | submgmid 42118 | Every magma is trivially a submagma of itself. (Contributed by AV, 26-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑀) ⇒ ⊢ (𝑀 ∈ Mgm → 𝐵 ∈ (SubMgm‘𝑀)) | ||
Theorem | submgmcl 42119 | Submagmas are closed under the monoid operation. (Contributed by AV, 26-Feb-2020.) |
⊢ + = (+g‘𝑀) ⇒ ⊢ ((𝑆 ∈ (SubMgm‘𝑀) ∧ 𝑋 ∈ 𝑆 ∧ 𝑌 ∈ 𝑆) → (𝑋 + 𝑌) ∈ 𝑆) | ||
Theorem | submgmmgm 42120 | Submagmas are themselves magmas under the given operation. (Contributed by AV, 26-Feb-2020.) |
⊢ 𝐻 = (𝑀 ↾s 𝑆) ⇒ ⊢ (𝑆 ∈ (SubMgm‘𝑀) → 𝐻 ∈ Mgm) | ||
Theorem | submgmbas 42121 | The base set of a submagma. (Contributed by AV, 26-Feb-2020.) |
⊢ 𝐻 = (𝑀 ↾s 𝑆) ⇒ ⊢ (𝑆 ∈ (SubMgm‘𝑀) → 𝑆 = (Base‘𝐻)) | ||
Theorem | subsubmgm 42122 | A submagma of a submagma is a submagma. (Contributed by AV, 26-Feb-2020.) |
⊢ 𝐻 = (𝐺 ↾s 𝑆) ⇒ ⊢ (𝑆 ∈ (SubMgm‘𝐺) → (𝐴 ∈ (SubMgm‘𝐻) ↔ (𝐴 ∈ (SubMgm‘𝐺) ∧ 𝐴 ⊆ 𝑆))) | ||
Theorem | resmgmhm 42123 | Restriction of a magma homomorphism to a submagma is a homomorphism. (Contributed by AV, 26-Feb-2020.) |
⊢ 𝑈 = (𝑆 ↾s 𝑋) ⇒ ⊢ ((𝐹 ∈ (𝑆 MgmHom 𝑇) ∧ 𝑋 ∈ (SubMgm‘𝑆)) → (𝐹 ↾ 𝑋) ∈ (𝑈 MgmHom 𝑇)) | ||
Theorem | resmgmhm2 42124 | One direction of resmgmhm2b 42125. (Contributed by AV, 26-Feb-2020.) |
⊢ 𝑈 = (𝑇 ↾s 𝑋) ⇒ ⊢ ((𝐹 ∈ (𝑆 MgmHom 𝑈) ∧ 𝑋 ∈ (SubMgm‘𝑇)) → 𝐹 ∈ (𝑆 MgmHom 𝑇)) | ||
Theorem | resmgmhm2b 42125 | Restriction of the codomain of a homomorphism. (Contributed by AV, 26-Feb-2020.) |
⊢ 𝑈 = (𝑇 ↾s 𝑋) ⇒ ⊢ ((𝑋 ∈ (SubMgm‘𝑇) ∧ ran 𝐹 ⊆ 𝑋) → (𝐹 ∈ (𝑆 MgmHom 𝑇) ↔ 𝐹 ∈ (𝑆 MgmHom 𝑈))) | ||
Theorem | mgmhmco 42126 | The composition of magma homomorphisms is a homomorphism. (Contributed by AV, 27-Feb-2020.) |
⊢ ((𝐹 ∈ (𝑇 MgmHom 𝑈) ∧ 𝐺 ∈ (𝑆 MgmHom 𝑇)) → (𝐹 ∘ 𝐺) ∈ (𝑆 MgmHom 𝑈)) | ||
Theorem | mgmhmima 42127 | The homomorphic image of a submagma is a submagma. (Contributed by AV, 27-Feb-2020.) |
⊢ ((𝐹 ∈ (𝑀 MgmHom 𝑁) ∧ 𝑋 ∈ (SubMgm‘𝑀)) → (𝐹 “ 𝑋) ∈ (SubMgm‘𝑁)) | ||
Theorem | mgmhmeql 42128 | The equalizer of two magma homomorphisms is a submagma. (Contributed by AV, 27-Feb-2020.) |
⊢ ((𝐹 ∈ (𝑆 MgmHom 𝑇) ∧ 𝐺 ∈ (𝑆 MgmHom 𝑇)) → dom (𝐹 ∩ 𝐺) ∈ (SubMgm‘𝑆)) | ||
Theorem | submgmacs 42129 | Submagmas are an algebraic closure system. (Contributed by AV, 27-Feb-2020.) |
⊢ 𝐵 = (Base‘𝐺) ⇒ ⊢ (𝐺 ∈ Mgm → (SubMgm‘𝐺) ∈ (ACS‘𝐵)) | ||
Theorem | ismhm0 42130 | Property of a monoid homomorphism, expressed by a magma homomorphism. (Contributed by AV, 17-Apr-2020.) |
⊢ 𝐵 = (Base‘𝑆) & ⊢ 𝐶 = (Base‘𝑇) & ⊢ + = (+g‘𝑆) & ⊢ ⨣ = (+g‘𝑇) & ⊢ 0 = (0g‘𝑆) & ⊢ 𝑌 = (0g‘𝑇) ⇒ ⊢ (𝐹 ∈ (𝑆 MndHom 𝑇) ↔ ((𝑆 ∈ Mnd ∧ 𝑇 ∈ Mnd) ∧ (𝐹 ∈ (𝑆 MgmHom 𝑇) ∧ (𝐹‘ 0 ) = 𝑌))) | ||
Theorem | mhmismgmhm 42131 | Each monoid homomorphism is a magma homomorphism. (Contributed by AV, 29-Feb-2020.) |
⊢ (𝐹 ∈ (𝑅 MndHom 𝑆) → 𝐹 ∈ (𝑅 MgmHom 𝑆)) | ||
Theorem | opmpt2ismgm 42132* | A structure with a group addition operation in maps-to notation is a magma if the operation value is contained in the base set. (Contributed by AV, 16-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ (+g‘𝑀) = (𝑥 ∈ 𝐵, 𝑦 ∈ 𝐵 ↦ 𝐶) & ⊢ (𝜑 → 𝐵 ≠ ∅) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝐵 ∧ 𝑦 ∈ 𝐵)) → 𝐶 ∈ 𝐵) ⇒ ⊢ (𝜑 → 𝑀 ∈ Mgm) | ||
Theorem | copissgrp 42133* | A structure with a constant group addition operation is a semigroup if the constant is contained in the base set. (Contributed by AV, 16-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ (+g‘𝑀) = (𝑥 ∈ 𝐵, 𝑦 ∈ 𝐵 ↦ 𝐶) & ⊢ (𝜑 → 𝐵 ≠ ∅) & ⊢ (𝜑 → 𝐶 ∈ 𝐵) ⇒ ⊢ (𝜑 → 𝑀 ∈ SGrp) | ||
Theorem | copisnmnd 42134* | A structure with a constant group addition operation and at least two elements is not a monoid. (Contributed by AV, 16-Feb-2020.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ (+g‘𝑀) = (𝑥 ∈ 𝐵, 𝑦 ∈ 𝐵 ↦ 𝐶) & ⊢ (𝜑 → 𝐶 ∈ 𝐵) & ⊢ (𝜑 → 1 < (#‘𝐵)) ⇒ ⊢ (𝜑 → 𝑀 ∉ Mnd) | ||
Theorem | 0nodd 42135* | 0 is not an odd integer. (Contributed by AV, 3-Feb-2020.) |
⊢ 𝑂 = {𝑧 ∈ ℤ ∣ ∃𝑥 ∈ ℤ 𝑧 = ((2 · 𝑥) + 1)} ⇒ ⊢ 0 ∉ 𝑂 | ||
Theorem | 1odd 42136* | 1 is an odd integer. (Contributed by AV, 3-Feb-2020.) |
⊢ 𝑂 = {𝑧 ∈ ℤ ∣ ∃𝑥 ∈ ℤ 𝑧 = ((2 · 𝑥) + 1)} ⇒ ⊢ 1 ∈ 𝑂 | ||
Theorem | 2nodd 42137* | 2 is not an odd integer. (Contributed by AV, 3-Feb-2020.) |
⊢ 𝑂 = {𝑧 ∈ ℤ ∣ ∃𝑥 ∈ ℤ 𝑧 = ((2 · 𝑥) + 1)} ⇒ ⊢ 2 ∉ 𝑂 | ||
Theorem | oddibas 42138* | Lemma 1 for oddinmgm 42140: The base set of M is the set of all odd integers. (Contributed by AV, 3-Feb-2020.) |
⊢ 𝑂 = {𝑧 ∈ ℤ ∣ ∃𝑥 ∈ ℤ 𝑧 = ((2 · 𝑥) + 1)} & ⊢ 𝑀 = (ℂfld ↾s 𝑂) ⇒ ⊢ 𝑂 = (Base‘𝑀) | ||
Theorem | oddiadd 42139* | Lemma 2 for oddinmgm 42140: The group addition operation of M is the addition of complex numbers. (Contributed by AV, 3-Feb-2020.) |
⊢ 𝑂 = {𝑧 ∈ ℤ ∣ ∃𝑥 ∈ ℤ 𝑧 = ((2 · 𝑥) + 1)} & ⊢ 𝑀 = (ℂfld ↾s 𝑂) ⇒ ⊢ + = (+g‘𝑀) | ||
Theorem | oddinmgm 42140* | The structure of all odd integers together with the addition of complex numbers is not a magma. Remark: the structure of the complementary subset of the set of integers, the even integers, is a magma, actually an abelian group, see 2zrngaabl 42269, and even a non-unital ring, see 2zrng 42260. (Contributed by AV, 3-Feb-2020.) |
⊢ 𝑂 = {𝑧 ∈ ℤ ∣ ∃𝑥 ∈ ℤ 𝑧 = ((2 · 𝑥) + 1)} & ⊢ 𝑀 = (ℂfld ↾s 𝑂) ⇒ ⊢ 𝑀 ∉ Mgm | ||
Theorem | nnsgrpmgm 42141 | The structure of positive integers together with the addition of complex numbers is a magma. (Contributed by AV, 4-Feb-2020.) |
⊢ 𝑀 = (ℂfld ↾s ℕ) ⇒ ⊢ 𝑀 ∈ Mgm | ||
Theorem | nnsgrp 42142 | The structure of positive integers together with the addition of complex numbers is a semigroup. (Contributed by AV, 4-Feb-2020.) |
⊢ 𝑀 = (ℂfld ↾s ℕ) ⇒ ⊢ 𝑀 ∈ SGrp | ||
Theorem | nnsgrpnmnd 42143 | The structure of positive integers together with the addition of complex numbers is not a monoid. (Contributed by AV, 4-Feb-2020.) |
⊢ 𝑀 = (ℂfld ↾s ℕ) ⇒ ⊢ 𝑀 ∉ Mnd | ||
With df-mpt2 6695, binary operations are defined by a rule, and with df-ov 6693, the value of a binary operation applied to two operands can be expressed. In both cases, the two operands can belong to different sets, and the result can be an element of a third set. However, according to Wikipedia "Binary operation", see https://en.wikipedia.org/wiki/Binary_operation (19-Jan-2020), "... a binary operation on a set 𝑆 is a mapping of the elements of the Cartesian product 𝑆 × 𝑆 to S: 𝑓:𝑆 × 𝑆⟶𝑆. Because the result of performing the operation on a pair of elements of S is again an element of S, the operation is called a closed binary operation on S (or sometimes expressed as having the property of closure).". To distinguish this more restrictive definition (in Wikipedia and most of the literature) from the general case, we call binary operations mapping the elements of the Cartesian product 𝑆 × 𝑆 internal binary operations, see df-intop 42160. If, in addition, the result is also contained in the set 𝑆, the operation is called closed internal binary operation, see df-clintop 42161. Therefore, a "binary operation on a set 𝑆" according to Wikipedia is a "closed internal binary operation" in our terminology. If the sets are different, the operation is explicitly called external binary operation (see Wikipedia https://en.wikipedia.org/wiki/Binary_operation#External_binary_operations ). Taking a step back, we define "laws" applicable for "binary operations" (which even need not to be functions), according to the definition in [Hall] p. 1 and [BourbakiAlg1] p. 1, p. 4 and p. 7. These laws are used, on the one hand, to specialize internal binary operations (see df-clintop 42161 and df-assintop 42162), and on the other hand to define the common algebraic structures like magmas, groups, rings, etc. Internal binary operations, which obey these laws, are defined afterwards. Notice that in [BourbakiAlg1] p. 1, p. 4 and p. 7, these operations are called "laws" by themselves. In the following, an alternate definition df-cllaw 42147 for an internal binary operation is provided, which does not require function-ness, but only closure. Therefore, this definition could be used as binary operation (Slot 2) defined for a magma as extensible structure, see mgmplusgiopALT 42155, or for an alternate definition df-mgm2 42180 for a magma as extensible structure. Similar results are obtained for an associative operation (defining semigroups). | ||
In this subsection, the "laws" applicable for "binary operations" according to the definition in [Hall] p. 1 and [BourbakiAlg1] p. 1, p. 4 and p. 7 are defined. These laws are called "internal laws" in [BourbakiAlg1] p. xxi. | ||
Syntax | ccllaw 42144 | Extend class notation for the closure law. |
class clLaw | ||
Syntax | casslaw 42145 | Extend class notation for the associative law. |
class assLaw | ||
Syntax | ccomlaw 42146 | Extend class notation for the commutative law. |
class comLaw | ||
Definition | df-cllaw 42147* | The closure law for binary operations, see definitions of laws A0. and M0. in section 1.1 of [Hall] p. 1, or definition 1 in [BourbakiAlg1] p. 1: the value of a binary operation applied to two operands of a given sets is an element of this set. By this definition, the closure law is expressed as binary relation: a binary operation is related to a set by clLaw if the closure law holds for this binary operation regarding this set. Note that the binary operation needs not to be a function. (Contributed by AV, 7-Jan-2020.) |
⊢ clLaw = {〈𝑜, 𝑚〉 ∣ ∀𝑥 ∈ 𝑚 ∀𝑦 ∈ 𝑚 (𝑥𝑜𝑦) ∈ 𝑚} | ||
Definition | df-comlaw 42148* | The commutative law for binary operations, see definitions of laws A2. and M2. in section 1.1 of [Hall] p. 1, or definition 8 in [BourbakiAlg1] p. 7: the value of a binary operation applied to two operands equals the value of a binary operation applied to the two operands in reversed order. By this definition, the commutative law is expressed as binary relation: a binary operation is related to a set by comLaw if the commutative law holds for this binary operation regarding this set. Note that the binary operation needs neither to be closed nor to be a function. (Contributed by AV, 7-Jan-2020.) |
⊢ comLaw = {〈𝑜, 𝑚〉 ∣ ∀𝑥 ∈ 𝑚 ∀𝑦 ∈ 𝑚 (𝑥𝑜𝑦) = (𝑦𝑜𝑥)} | ||
Definition | df-asslaw 42149* | The associative law for binary operations, see definitions of laws A1. and M1. in section 1.1 of [Hall] p. 1, or definition 5 in [BourbakiAlg1] p. 4: the value of a binary operation applied the value of the binary operation applied to two operands and a third operand equals the value of the binary operation applied to the first operand and the value of the binary operation applied to the second and third operand. By this definition, the associative law is expressed as binary relation: a binary operation is related to a set by assLaw if the associative law holds for this binary operation regarding this set. Note that the binary operation needs neither to be closed nor to be a function. (Contributed by FL, 1-Nov-2009.) (Revised by AV, 13-Jan-2020.) |
⊢ assLaw = {〈𝑜, 𝑚〉 ∣ ∀𝑥 ∈ 𝑚 ∀𝑦 ∈ 𝑚 ∀𝑧 ∈ 𝑚 ((𝑥𝑜𝑦)𝑜𝑧) = (𝑥𝑜(𝑦𝑜𝑧))} | ||
Theorem | iscllaw 42150* | The predicate "is a closed operation". (Contributed by AV, 13-Jan-2020.) |
⊢ (( ⚬ ∈ 𝑉 ∧ 𝑀 ∈ 𝑊) → ( ⚬ clLaw 𝑀 ↔ ∀𝑥 ∈ 𝑀 ∀𝑦 ∈ 𝑀 (𝑥 ⚬ 𝑦) ∈ 𝑀)) | ||
Theorem | iscomlaw 42151* | The predicate "is a commutative operation". (Contributed by AV, 20-Jan-2020.) |
⊢ (( ⚬ ∈ 𝑉 ∧ 𝑀 ∈ 𝑊) → ( ⚬ comLaw 𝑀 ↔ ∀𝑥 ∈ 𝑀 ∀𝑦 ∈ 𝑀 (𝑥 ⚬ 𝑦) = (𝑦 ⚬ 𝑥))) | ||
Theorem | clcllaw 42152 | Closure of a closed operation. (Contributed by FL, 14-Sep-2010.) (Revised by AV, 21-Jan-2020.) |
⊢ (( ⚬ clLaw 𝑀 ∧ 𝑋 ∈ 𝑀 ∧ 𝑌 ∈ 𝑀) → (𝑋 ⚬ 𝑌) ∈ 𝑀) | ||
Theorem | isasslaw 42153* | The predicate "is an associative operation". (Contributed by FL, 1-Nov-2009.) (Revised by AV, 13-Jan-2020.) |
⊢ (( ⚬ ∈ 𝑉 ∧ 𝑀 ∈ 𝑊) → ( ⚬ assLaw 𝑀 ↔ ∀𝑥 ∈ 𝑀 ∀𝑦 ∈ 𝑀 ∀𝑧 ∈ 𝑀 ((𝑥 ⚬ 𝑦) ⚬ 𝑧) = (𝑥 ⚬ (𝑦 ⚬ 𝑧)))) | ||
Theorem | asslawass 42154* | Associativity of an associative operation. (Contributed by FL, 2-Nov-2009.) (Revised by AV, 21-Jan-2020.) |
⊢ ( ⚬ assLaw 𝑀 → ∀𝑥 ∈ 𝑀 ∀𝑦 ∈ 𝑀 ∀𝑧 ∈ 𝑀 ((𝑥 ⚬ 𝑦) ⚬ 𝑧) = (𝑥 ⚬ (𝑦 ⚬ 𝑧))) | ||
Theorem | mgmplusgiopALT 42155 | Slot 2 (group operation) of a magma as extensible structure is a closed operation on the base set. (Contributed by AV, 13-Jan-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝑀 ∈ Mgm → (+g‘𝑀) clLaw (Base‘𝑀)) | ||
Theorem | sgrpplusgaopALT 42156 | Slot 2 (group operation) of a semigroup as extensible structure is an associative operation on the base set. (Contributed by AV, 13-Jan-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝐺 ∈ SGrp → (+g‘𝐺) assLaw (Base‘𝐺)) | ||
In this subsection, "internal binary operations" obeying different laws are defined. | ||
Syntax | cintop 42157 | Extend class notation with class of internal (binary) operations for a set. |
class intOp | ||
Syntax | cclintop 42158 | Extend class notation with class of closed operations for a set. |
class clIntOp | ||
Syntax | cassintop 42159 | Extend class notation with class of associative operations for a set. |
class assIntOp | ||
Definition | df-intop 42160* | Function mapping a set to the class of all internal (binary) operations for this set. (Contributed by AV, 20-Jan-2020.) |
⊢ intOp = (𝑚 ∈ V, 𝑛 ∈ V ↦ (𝑛 ↑𝑚 (𝑚 × 𝑚))) | ||
Definition | df-clintop 42161 | Function mapping a set to the class of all closed (internal binary) operations for this set, see definition in section 1.2 of [Hall] p. 2, definition in section I.1 of [Bruck] p. 1, or definition 1 in [BourbakiAlg1] p. 1, where it is called "a law of composition". (Contributed by AV, 20-Jan-2020.) |
⊢ clIntOp = (𝑚 ∈ V ↦ (𝑚 intOp 𝑚)) | ||
Definition | df-assintop 42162* | Function mapping a set to the class of all associative (closed internal binary) operations for this set, see definition 5 in [BourbakiAlg1] p. 4, where it is called "an associative law of composition". (Contributed by AV, 20-Jan-2020.) |
⊢ assIntOp = (𝑚 ∈ V ↦ {𝑜 ∈ ( clIntOp ‘𝑚) ∣ 𝑜 assLaw 𝑚}) | ||
Theorem | intopval 42163 | The internal (binary) operations for a set. (Contributed by AV, 20-Jan-2020.) |
⊢ ((𝑀 ∈ 𝑉 ∧ 𝑁 ∈ 𝑊) → (𝑀 intOp 𝑁) = (𝑁 ↑𝑚 (𝑀 × 𝑀))) | ||
Theorem | intop 42164 | An internal (binary) operation for a set. (Contributed by AV, 20-Jan-2020.) |
⊢ ( ⚬ ∈ (𝑀 intOp 𝑁) → ⚬ :(𝑀 × 𝑀)⟶𝑁) | ||
Theorem | clintopval 42165 | The closed (internal binary) operations for a set. (Contributed by AV, 20-Jan-2020.) |
⊢ (𝑀 ∈ 𝑉 → ( clIntOp ‘𝑀) = (𝑀 ↑𝑚 (𝑀 × 𝑀))) | ||
Theorem | assintopval 42166* | The associative (closed internal binary) operations for a set. (Contributed by AV, 20-Jan-2020.) |
⊢ (𝑀 ∈ 𝑉 → ( assIntOp ‘𝑀) = {𝑜 ∈ ( clIntOp ‘𝑀) ∣ 𝑜 assLaw 𝑀}) | ||
Theorem | assintopmap 42167* | The associative (closed internal binary) operations for a set, expressed with set exponentiation. (Contributed by AV, 20-Jan-2020.) |
⊢ (𝑀 ∈ 𝑉 → ( assIntOp ‘𝑀) = {𝑜 ∈ (𝑀 ↑𝑚 (𝑀 × 𝑀)) ∣ 𝑜 assLaw 𝑀}) | ||
Theorem | isclintop 42168 | The predicate "is a closed (internal binary) operations for a set". (Contributed by FL, 2-Nov-2009.) (Revised by AV, 20-Jan-2020.) |
⊢ (𝑀 ∈ 𝑉 → ( ⚬ ∈ ( clIntOp ‘𝑀) ↔ ⚬ :(𝑀 × 𝑀)⟶𝑀)) | ||
Theorem | clintop 42169 | A closed (internal binary) operation for a set. (Contributed by AV, 20-Jan-2020.) |
⊢ ( ⚬ ∈ ( clIntOp ‘𝑀) → ⚬ :(𝑀 × 𝑀)⟶𝑀) | ||
Theorem | assintop 42170 | An associative (closed internal binary) operation for a set. (Contributed by AV, 20-Jan-2020.) |
⊢ ( ⚬ ∈ ( assIntOp ‘𝑀) → ( ⚬ :(𝑀 × 𝑀)⟶𝑀 ∧ ⚬ assLaw 𝑀)) | ||
Theorem | isassintop 42171* | The predicate "is an associative (closed internal binary) operations for a set". (Contributed by FL, 2-Nov-2009.) (Revised by AV, 20-Jan-2020.) |
⊢ (𝑀 ∈ 𝑉 → ( ⚬ ∈ ( assIntOp ‘𝑀) ↔ ( ⚬ :(𝑀 × 𝑀)⟶𝑀 ∧ ∀𝑥 ∈ 𝑀 ∀𝑦 ∈ 𝑀 ∀𝑧 ∈ 𝑀 ((𝑥 ⚬ 𝑦) ⚬ 𝑧) = (𝑥 ⚬ (𝑦 ⚬ 𝑧))))) | ||
Theorem | clintopcllaw 42172 | The closure law holds for a closed (internal binary) operation for a set. (Contributed by AV, 20-Jan-2020.) |
⊢ ( ⚬ ∈ ( clIntOp ‘𝑀) → ⚬ clLaw 𝑀) | ||
Theorem | assintopcllaw 42173 | The closure low holds for an associative (closed internal binary) operation for a set. (Contributed by FL, 2-Nov-2009.) (Revised by AV, 20-Jan-2020.) |
⊢ ( ⚬ ∈ ( assIntOp ‘𝑀) → ⚬ clLaw 𝑀) | ||
Theorem | assintopasslaw 42174 | The associative low holds for a associative (closed internal binary) operation for a set. (Contributed by FL, 2-Nov-2009.) (Revised by AV, 20-Jan-2020.) |
⊢ ( ⚬ ∈ ( assIntOp ‘𝑀) → ⚬ assLaw 𝑀) | ||
Theorem | assintopass 42175* | An associative (closed internal binary) operation for a set is associative. (Contributed by FL, 2-Nov-2009.) (Revised by AV, 20-Jan-2020.) |
⊢ ( ⚬ ∈ ( assIntOp ‘𝑀) → ∀𝑥 ∈ 𝑀 ∀𝑦 ∈ 𝑀 ∀𝑧 ∈ 𝑀 ((𝑥 ⚬ 𝑦) ⚬ 𝑧) = (𝑥 ⚬ (𝑦 ⚬ 𝑧))) | ||
Syntax | cmgm2 42176 | Extend class notation with class of all magmas. |
class MgmALT | ||
Syntax | ccmgm2 42177 | Extend class notation with class of all commutative magmas. |
class CMgmALT | ||
Syntax | csgrp2 42178 | Extend class notation with class of all semigroups. |
class SGrpALT | ||
Syntax | ccsgrp2 42179 | Extend class notation with class of all commutative semigroups. |
class CSGrpALT | ||
Definition | df-mgm2 42180 | A magma is a set equipped with a closed operation. Definition 1 of [BourbakiAlg1] p. 1, or definition of a groupoid in section I.1 of [Bruck] p. 1. Note: The term "groupoid" is now widely used to refer to other objects: (small) categories all of whose morphisms are invertible, or groups with a partial function replacing the binary operation. Therefore, we will only use the term "magma" for the present notion in set.mm. (Contributed by AV, 6-Jan-2020.) |
⊢ MgmALT = {𝑚 ∣ (+g‘𝑚) clLaw (Base‘𝑚)} | ||
Definition | df-cmgm2 42181 | A commutative magma is a magma with a commutative operation. Definition 8 of [BourbakiAlg1] p. 7. (Contributed by AV, 20-Jan-2020.) |
⊢ CMgmALT = {𝑚 ∈ MgmALT ∣ (+g‘𝑚) comLaw (Base‘𝑚)} | ||
Definition | df-sgrp2 42182 | A semigroup is a magma with an associative operation. Definition in section II.1 of [Bruck] p. 23, or of an "associative magma" in definition 5 of [BourbakiAlg1] p. 4, or of a semi-group in section 1.3 of [Hall] p. 7. (Contributed by AV, 6-Jan-2020.) |
⊢ SGrpALT = {𝑔 ∈ MgmALT ∣ (+g‘𝑔) assLaw (Base‘𝑔)} | ||
Definition | df-csgrp2 42183 | A commutative semigroup is a semigroup with a commutative operation. (Contributed by AV, 20-Jan-2020.) |
⊢ CSGrpALT = {𝑔 ∈ SGrpALT ∣ (+g‘𝑔) comLaw (Base‘𝑔)} | ||
Theorem | ismgmALT 42184 | The predicate "is a magma." (Contributed by AV, 16-Jan-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ ⚬ = (+g‘𝑀) ⇒ ⊢ (𝑀 ∈ 𝑉 → (𝑀 ∈ MgmALT ↔ ⚬ clLaw 𝐵)) | ||
Theorem | iscmgmALT 42185 | The predicate "is a commutative magma." (Contributed by AV, 20-Jan-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ ⚬ = (+g‘𝑀) ⇒ ⊢ (𝑀 ∈ CMgmALT ↔ (𝑀 ∈ MgmALT ∧ ⚬ comLaw 𝐵)) | ||
Theorem | issgrpALT 42186 | The predicate "is a semigroup." (Contributed by AV, 16-Jan-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ ⚬ = (+g‘𝑀) ⇒ ⊢ (𝑀 ∈ SGrpALT ↔ (𝑀 ∈ MgmALT ∧ ⚬ assLaw 𝐵)) | ||
Theorem | iscsgrpALT 42187 | The predicate "is a commutative semigroup." (Contributed by AV, 20-Jan-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ 𝐵 = (Base‘𝑀) & ⊢ ⚬ = (+g‘𝑀) ⇒ ⊢ (𝑀 ∈ CSGrpALT ↔ (𝑀 ∈ SGrpALT ∧ ⚬ comLaw 𝐵)) | ||
Theorem | mgm2mgm 42188 | Equivalence of the two definitions of a magma. (Contributed by AV, 16-Jan-2020.) |
⊢ (𝑀 ∈ MgmALT ↔ 𝑀 ∈ Mgm) | ||
Theorem | sgrp2sgrp 42189 | Equivalence of the two definitions of a semigroup. (Contributed by AV, 16-Jan-2020.) |
⊢ (𝑀 ∈ SGrpALT ↔ 𝑀 ∈ SGrp) | ||
Theorem | idfusubc0 42190* | The identity functor for a subcategory is an "inclusion functor" from the subcategory into its supercategory. (Contributed by AV, 29-Mar-2020.) |
⊢ 𝑆 = (𝐶 ↾cat 𝐽) & ⊢ 𝐼 = (idfunc‘𝑆) & ⊢ 𝐵 = (Base‘𝑆) ⇒ ⊢ (𝐽 ∈ (Subcat‘𝐶) → 𝐼 = 〈( I ↾ 𝐵), (𝑥 ∈ 𝐵, 𝑦 ∈ 𝐵 ↦ ( I ↾ (𝑥(Hom ‘𝑆)𝑦)))〉) | ||
Theorem | idfusubc 42191* | The identity functor for a subcategory is an "inclusion functor" from the subcategory into its supercategory. (Contributed by AV, 29-Mar-2020.) |
⊢ 𝑆 = (𝐶 ↾cat 𝐽) & ⊢ 𝐼 = (idfunc‘𝑆) & ⊢ 𝐵 = (Base‘𝑆) ⇒ ⊢ (𝐽 ∈ (Subcat‘𝐶) → 𝐼 = 〈( I ↾ 𝐵), (𝑥 ∈ 𝐵, 𝑦 ∈ 𝐵 ↦ ( I ↾ (𝑥𝐽𝑦)))〉) | ||
Theorem | inclfusubc 42192* | The "inclusion functor" from a subcategory of a category into the category itself. (Contributed by AV, 30-Mar-2020.) |
⊢ (𝜑 → 𝐽 ∈ (Subcat‘𝐶)) & ⊢ 𝑆 = (𝐶 ↾cat 𝐽) & ⊢ 𝐵 = (Base‘𝑆) & ⊢ (𝜑 → 𝐹 = ( I ↾ 𝐵)) & ⊢ (𝜑 → 𝐺 = (𝑥 ∈ 𝐵, 𝑦 ∈ 𝐵 ↦ ( I ↾ (𝑥𝐽𝑦)))) ⇒ ⊢ (𝜑 → 𝐹(𝑆 Func 𝐶)𝐺) | ||
Theorem | lmod0rng 42193 | If the scalar ring of a module is the zero ring, the module is the zero module, i.e. the base set of the module is the singleton consisting of the identity element only. (Contributed by AV, 17-Apr-2019.) |
⊢ ((𝑀 ∈ LMod ∧ ¬ (Scalar‘𝑀) ∈ NzRing) → (Base‘𝑀) = {(0g‘𝑀)}) | ||
Theorem | nzrneg1ne0 42194 | The additive inverse of the 1 in a nonzero ring is not zero ( -1 =/= 0 ). (Contributed by AV, 29-Apr-2019.) |
⊢ (𝑅 ∈ NzRing → ((invg‘𝑅)‘(1r‘𝑅)) ≠ (0g‘𝑅)) | ||
Theorem | 0ringdif 42195 | A zero ring is a ring which is not a nonzero ring. (Contributed by AV, 17-Apr-2020.) |
⊢ 𝐵 = (Base‘𝑅) & ⊢ 0 = (0g‘𝑅) ⇒ ⊢ (𝑅 ∈ (Ring ∖ NzRing) ↔ (𝑅 ∈ Ring ∧ 𝐵 = { 0 })) | ||
Theorem | 0ringbas 42196 | The base set of a zero ring, a ring which is not a nonzero ring, is the singleton of the zero element. (Contributed by AV, 17-Apr-2020.) |
⊢ 𝐵 = (Base‘𝑅) & ⊢ 0 = (0g‘𝑅) ⇒ ⊢ (𝑅 ∈ (Ring ∖ NzRing) → 𝐵 = { 0 }) | ||
Theorem | 0ring1eq0 42197 | In a zero ring, a ring which is not a nonzero ring, the unit equals the zero element. (Contributed by AV, 17-Apr-2020.) |
⊢ 𝐵 = (Base‘𝑅) & ⊢ 0 = (0g‘𝑅) & ⊢ 1 = (1r‘𝑅) ⇒ ⊢ (𝑅 ∈ (Ring ∖ NzRing) → 1 = 0 ) | ||
Theorem | nrhmzr 42198 | There is no ring homomorphism from the zero ring into a nonzero ring. (Contributed by AV, 18-Apr-2020.) |
⊢ ((𝑍 ∈ (Ring ∖ NzRing) ∧ 𝑅 ∈ NzRing) → (𝑍 RingHom 𝑅) = ∅) | ||
According to Wikipedia, "... in abstract algebra, a rng (or pseudo-ring or non-unital ring) is an algebraic structure satisfying the same properties as a [unital] ring, without assuming the existence of a multiplicative identity. The term "rng" (pronounced rung) is meant to suggest that it is a "ring" without "i", i.e. without the requirement for an "identity element"." (see https://en.wikipedia.org/wiki/Rng_(algebra), 6-Jan-2020). | ||
Syntax | crng 42199 | Extend class notation with class of all non-unital rings. |
class Rng | ||
Definition | df-rng0 42200* | Define class of all (non-unital) rings. A non-unital ring (or rng, or pseudoring) is a set equipped with two everywhere-defined internal operations, whose first one is an additive abelian group operation and the second one is a multiplicative semigroup operation, and where the addition is left- and right-distributive for the multiplication. Definition of a pseudo-ring in section I.8.1 of [BourbakiAlg1] p. 93 or the definition of a ring in part Preliminaries of [Roman] p. 18. As almost always in mathematics, "non-unital" means "not necessarily unital". Therefore, by talking about a ring (in general) or a non-unital ring the "unital" case is always included. In contrast to a unital ring, the commutativity of addition must be postulated and cannot be proven from the other conditions. (Contributed by AV, 6-Jan-2020.) |
⊢ Rng = {𝑓 ∈ Abel ∣ ((mulGrp‘𝑓) ∈ SGrp ∧ [(Base‘𝑓) / 𝑏][(+g‘𝑓) / 𝑝][(.r‘𝑓) / 𝑡]∀𝑥 ∈ 𝑏 ∀𝑦 ∈ 𝑏 ∀𝑧 ∈ 𝑏 ((𝑥𝑡(𝑦𝑝𝑧)) = ((𝑥𝑡𝑦)𝑝(𝑥𝑡𝑧)) ∧ ((𝑥𝑝𝑦)𝑡𝑧) = ((𝑥𝑡𝑧)𝑝(𝑦𝑡𝑧))))} |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |