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

Theorem ismnd 12686
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 12690), whose operation is associative (so, a semigroup, see also mndass 12691) and has a two-sided neutral element (see mndid 12692). (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 12685 . 2 (𝐺 ∈ Mnd ↔ (𝐺 ∈ Smgrp ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)))
4 rexm 3520 . . . . 5 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → ∃𝑒 𝑒𝐵)
5 eleq1w 2236 . . . . . 6 (𝑒 = 𝑤 → (𝑒𝐵𝑤𝐵))
65cbvexv 1916 . . . . 5 (∃𝑒 𝑒𝐵 ↔ ∃𝑤 𝑤𝐵)
74, 6sylib 122 . . . 4 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → ∃𝑤 𝑤𝐵)
81basmex 12487 . . . . 5 (𝑤𝐵𝐺 ∈ V)
98exlimiv 1596 . . . 4 (∃𝑤 𝑤𝐵𝐺 ∈ V)
101, 2issgrpv 12676 . . . 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 1353  wex 1490  wcel 2146  wral 2453  wrex 2454  Vcvv 2735  cfv 5208  (class class class)co 5865  Basecbs 12429  +gcplusg 12493  Smgrpcsgrp 12673  Mndcmnd 12683
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 709  ax-5 1445  ax-7 1446  ax-gen 1447  ax-ie1 1491  ax-ie2 1492  ax-8 1502  ax-10 1503  ax-11 1504  ax-i12 1505  ax-bndl 1507  ax-4 1508  ax-17 1524  ax-i9 1528  ax-ial 1532  ax-i5r 1533  ax-13 2148  ax-14 2149  ax-ext 2157  ax-sep 4116  ax-pow 4169  ax-pr 4203  ax-un 4427  ax-cnex 7877  ax-resscn 7878  ax-1re 7880  ax-addrcl 7883
This theorem depends on definitions:  df-bi 117  df-3an 980  df-tru 1356  df-nf 1459  df-sb 1761  df-eu 2027  df-mo 2028  df-clab 2162  df-cleq 2168  df-clel 2171  df-nfc 2306  df-ral 2458  df-rex 2459  df-rab 2462  df-v 2737  df-sbc 2961  df-un 3131  df-in 3133  df-ss 3140  df-pw 3574  df-sn 3595  df-pr 3596  df-op 3598  df-uni 3806  df-int 3841  df-br 3999  df-opab 4060  df-mpt 4061  df-id 4287  df-xp 4626  df-rel 4627  df-cnv 4628  df-co 4629  df-dm 4630  df-rn 4631  df-res 4632  df-iota 5170  df-fun 5210  df-fn 5211  df-fv 5216  df-ov 5868  df-inn 8893  df-2 8951  df-ndx 12432  df-slot 12433  df-base 12435  df-plusg 12506  df-mgm 12641  df-sgrp 12674  df-mnd 12684
This theorem is referenced by:  mndid  12692  ismndd  12704  mndpropd  12707  mhmmnd  12841
  Copyright terms: Public domain W3C validator