ILE Home Intuitionistic Logic Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  ILE Home  >  Th. List  >  ismnd GIF version

Theorem ismnd 12655
Description: The predicate "is a monoid". This is the defining theorem of a monoid by showing that a set is a monoid if and only if it is a set equipped with a closed, everywhere defined internal operation (so, a magma, see mndcl 12659), whose operation is associative (so, a semigroup, see also mndass 12660) and has a two-sided neutral element (see mndid 12661). (Contributed by Mario Carneiro, 6-Jan-2015.) (Revised by AV, 1-Feb-2020.)
Hypotheses
Ref Expression
ismnd.b 𝐵 = (Base‘𝐺)
ismnd.p + = (+g𝐺)
Assertion
Ref Expression
ismnd (𝐺 ∈ Mnd ↔ (∀𝑎𝐵𝑏𝐵 ((𝑎 + 𝑏) ∈ 𝐵 ∧ ∀𝑐𝐵 ((𝑎 + 𝑏) + 𝑐) = (𝑎 + (𝑏 + 𝑐))) ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)))
Distinct variable groups:   𝐵,𝑎,𝑏,𝑐   𝐵,𝑒,𝑎   𝐺,𝑎,𝑏,𝑐   + ,𝑎,𝑒   + ,𝑏,𝑐
Allowed substitution hint:   𝐺(𝑒)

Proof of Theorem ismnd
Dummy variable 𝑤 is distinct from all other variables.
StepHypRef Expression
1 ismnd.b . . 3 𝐵 = (Base‘𝐺)
2 ismnd.p . . 3 + = (+g𝐺)
31, 2ismnddef 12654 . 2 (𝐺 ∈ Mnd ↔ (𝐺 ∈ Smgrp ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)))
4 rexm 3514 . . . . 5 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → ∃𝑒 𝑒𝐵)
5 eleq1w 2231 . . . . . 6 (𝑒 = 𝑤 → (𝑒𝐵𝑤𝐵))
65cbvexv 1911 . . . . 5 (∃𝑒 𝑒𝐵 ↔ ∃𝑤 𝑤𝐵)
74, 6sylib 121 . . . 4 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → ∃𝑤 𝑤𝐵)
81basmex 12474 . . . . 5 (𝑤𝐵𝐺 ∈ V)
98exlimiv 1591 . . . 4 (∃𝑤 𝑤𝐵𝐺 ∈ V)
101, 2issgrpv 12645 . . . 4 (𝐺 ∈ V → (𝐺 ∈ Smgrp ↔ ∀𝑎𝐵𝑏𝐵 ((𝑎 + 𝑏) ∈ 𝐵 ∧ ∀𝑐𝐵 ((𝑎 + 𝑏) + 𝑐) = (𝑎 + (𝑏 + 𝑐)))))
117, 9, 103syl 17 . . 3 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → (𝐺 ∈ Smgrp ↔ ∀𝑎𝐵𝑏𝐵 ((𝑎 + 𝑏) ∈ 𝐵 ∧ ∀𝑐𝐵 ((𝑎 + 𝑏) + 𝑐) = (𝑎 + (𝑏 + 𝑐)))))
1211pm5.32ri 452 . 2 ((𝐺 ∈ Smgrp ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)) ↔ (∀𝑎𝐵𝑏𝐵 ((𝑎 + 𝑏) ∈ 𝐵 ∧ ∀𝑐𝐵 ((𝑎 + 𝑏) + 𝑐) = (𝑎 + (𝑏 + 𝑐))) ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)))
133, 12bitri 183 1 (𝐺 ∈ Mnd ↔ (∀𝑎𝐵𝑏𝐵 ((𝑎 + 𝑏) ∈ 𝐵 ∧ ∀𝑐𝐵 ((𝑎 + 𝑏) + 𝑐) = (𝑎 + (𝑏 + 𝑐))) ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)))
Colors of variables: wff set class
Syntax hints:  wa 103  wb 104   = wceq 1348  wex 1485  wcel 2141  wral 2448  wrex 2449  Vcvv 2730  cfv 5198  (class class class)co 5853  Basecbs 12416  +gcplusg 12480  Smgrpcsgrp 12642  Mndcmnd 12652
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-ia1 105  ax-ia2 106  ax-ia3 107  ax-io 704  ax-5 1440  ax-7 1441  ax-gen 1442  ax-ie1 1486  ax-ie2 1487  ax-8 1497  ax-10 1498  ax-11 1499  ax-i12 1500  ax-bndl 1502  ax-4 1503  ax-17 1519  ax-i9 1523  ax-ial 1527  ax-i5r 1528  ax-13 2143  ax-14 2144  ax-ext 2152  ax-sep 4107  ax-pow 4160  ax-pr 4194  ax-un 4418  ax-cnex 7865  ax-resscn 7866  ax-1re 7868  ax-addrcl 7871
This theorem depends on definitions:  df-bi 116  df-3an 975  df-tru 1351  df-nf 1454  df-sb 1756  df-eu 2022  df-mo 2023  df-clab 2157  df-cleq 2163  df-clel 2166  df-nfc 2301  df-ral 2453  df-rex 2454  df-rab 2457  df-v 2732  df-sbc 2956  df-un 3125  df-in 3127  df-ss 3134  df-pw 3568  df-sn 3589  df-pr 3590  df-op 3592  df-uni 3797  df-int 3832  df-br 3990  df-opab 4051  df-mpt 4052  df-id 4278  df-xp 4617  df-rel 4618  df-cnv 4619  df-co 4620  df-dm 4621  df-rn 4622  df-res 4623  df-iota 5160  df-fun 5200  df-fn 5201  df-fv 5206  df-ov 5856  df-inn 8879  df-2 8937  df-ndx 12419  df-slot 12420  df-base 12422  df-plusg 12493  df-mgm 12610  df-sgrp 12643  df-mnd 12653
This theorem is referenced by:  mndid  12661  ismndd  12673  mndpropd  12676
  Copyright terms: Public domain W3C validator