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

Theorem ismnd 12895
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 12899), whose operation is associative (so, a semigroup, see also mndass 12900) and has a two-sided neutral element (see mndid 12901). (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 12894 . 2 (𝐺 ∈ Mnd ↔ (𝐺 ∈ Smgrp ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)))
4 rexm 3537 . . . . 5 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → ∃𝑒 𝑒𝐵)
5 eleq1w 2250 . . . . . 6 (𝑒 = 𝑤 → (𝑒𝐵𝑤𝐵))
65cbvexv 1930 . . . . 5 (∃𝑒 𝑒𝐵 ↔ ∃𝑤 𝑤𝐵)
74, 6sylib 122 . . . 4 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → ∃𝑤 𝑤𝐵)
81basmex 12574 . . . . 5 (𝑤𝐵𝐺 ∈ V)
98exlimiv 1609 . . . 4 (∃𝑤 𝑤𝐵𝐺 ∈ V)
101, 2issgrpv 12882 . . . 4 (𝐺 ∈ V → (𝐺 ∈ Smgrp ↔ ∀𝑎𝐵𝑏𝐵 ((𝑎 + 𝑏) ∈ 𝐵 ∧ ∀𝑐𝐵 ((𝑎 + 𝑏) + 𝑐) = (𝑎 + (𝑏 + 𝑐)))))
117, 9, 103syl 17 . . 3 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → (𝐺 ∈ Smgrp ↔ ∀𝑎𝐵𝑏𝐵 ((𝑎 + 𝑏) ∈ 𝐵 ∧ ∀𝑐𝐵 ((𝑎 + 𝑏) + 𝑐) = (𝑎 + (𝑏 + 𝑐)))))
1211pm5.32ri 455 . 2 ((𝐺 ∈ Smgrp ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)) ↔ (∀𝑎𝐵𝑏𝐵 ((𝑎 + 𝑏) ∈ 𝐵 ∧ ∀𝑐𝐵 ((𝑎 + 𝑏) + 𝑐) = (𝑎 + (𝑏 + 𝑐))) ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)))
133, 12bitri 184 1 (𝐺 ∈ Mnd ↔ (∀𝑎𝐵𝑏𝐵 ((𝑎 + 𝑏) ∈ 𝐵 ∧ ∀𝑐𝐵 ((𝑎 + 𝑏) + 𝑐) = (𝑎 + (𝑏 + 𝑐))) ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)))
Colors of variables: wff set class
Syntax hints:  wa 104  wb 105   = wceq 1364  wex 1503  wcel 2160  wral 2468  wrex 2469  Vcvv 2752  cfv 5235  (class class class)co 5897  Basecbs 12515  +gcplusg 12592  Smgrpcsgrp 12879  Mndcmnd 12892
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-ia1 106  ax-ia2 107  ax-ia3 108  ax-io 710  ax-5 1458  ax-7 1459  ax-gen 1460  ax-ie1 1504  ax-ie2 1505  ax-8 1515  ax-10 1516  ax-11 1517  ax-i12 1518  ax-bndl 1520  ax-4 1521  ax-17 1537  ax-i9 1541  ax-ial 1545  ax-i5r 1546  ax-13 2162  ax-14 2163  ax-ext 2171  ax-sep 4136  ax-pow 4192  ax-pr 4227  ax-un 4451  ax-cnex 7933  ax-resscn 7934  ax-1re 7936  ax-addrcl 7939
This theorem depends on definitions:  df-bi 117  df-3an 982  df-tru 1367  df-nf 1472  df-sb 1774  df-eu 2041  df-mo 2042  df-clab 2176  df-cleq 2182  df-clel 2185  df-nfc 2321  df-ral 2473  df-rex 2474  df-rab 2477  df-v 2754  df-sbc 2978  df-un 3148  df-in 3150  df-ss 3157  df-pw 3592  df-sn 3613  df-pr 3614  df-op 3616  df-uni 3825  df-int 3860  df-br 4019  df-opab 4080  df-mpt 4081  df-id 4311  df-xp 4650  df-rel 4651  df-cnv 4652  df-co 4653  df-dm 4654  df-rn 4655  df-res 4656  df-iota 5196  df-fun 5237  df-fn 5238  df-fv 5243  df-ov 5900  df-inn 8951  df-2 9009  df-ndx 12518  df-slot 12519  df-base 12521  df-plusg 12605  df-mgm 12835  df-sgrp 12880  df-mnd 12893
This theorem is referenced by:  mndid  12901  ismndd  12913  mndpropd  12916  mhmmnd  13073
  Copyright terms: Public domain W3C validator