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

Theorem ismnd 13000
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 13004), whose operation is associative (so, a semigroup, see also mndass 13005) and has a two-sided neutral element (see mndid 13006). (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 12999 . 2 (𝐺 ∈ Mnd ↔ (𝐺 ∈ Smgrp ∧ ∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎)))
4 rexm 3546 . . . . 5 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → ∃𝑒 𝑒𝐵)
5 eleq1w 2254 . . . . . 6 (𝑒 = 𝑤 → (𝑒𝐵𝑤𝐵))
65cbvexv 1930 . . . . 5 (∃𝑒 𝑒𝐵 ↔ ∃𝑤 𝑤𝐵)
74, 6sylib 122 . . . 4 (∃𝑒𝐵𝑎𝐵 ((𝑒 + 𝑎) = 𝑎 ∧ (𝑎 + 𝑒) = 𝑎) → ∃𝑤 𝑤𝐵)
81basmex 12677 . . . . 5 (𝑤𝐵𝐺 ∈ V)
98exlimiv 1609 . . . 4 (∃𝑤 𝑤𝐵𝐺 ∈ V)
101, 2issgrpv 12987 . . . 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 2164  wral 2472  wrex 2473  Vcvv 2760  cfv 5254  (class class class)co 5918  Basecbs 12618  +gcplusg 12695  Smgrpcsgrp 12984  Mndcmnd 12997
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 2166  ax-14 2167  ax-ext 2175  ax-sep 4147  ax-pow 4203  ax-pr 4238  ax-un 4464  ax-cnex 7963  ax-resscn 7964  ax-1re 7966  ax-addrcl 7969
This theorem depends on definitions:  df-bi 117  df-3an 982  df-tru 1367  df-nf 1472  df-sb 1774  df-eu 2045  df-mo 2046  df-clab 2180  df-cleq 2186  df-clel 2189  df-nfc 2325  df-ral 2477  df-rex 2478  df-rab 2481  df-v 2762  df-sbc 2986  df-un 3157  df-in 3159  df-ss 3166  df-pw 3603  df-sn 3624  df-pr 3625  df-op 3627  df-uni 3836  df-int 3871  df-br 4030  df-opab 4091  df-mpt 4092  df-id 4324  df-xp 4665  df-rel 4666  df-cnv 4667  df-co 4668  df-dm 4669  df-rn 4670  df-res 4671  df-iota 5215  df-fun 5256  df-fn 5257  df-fv 5262  df-ov 5921  df-inn 8983  df-2 9041  df-ndx 12621  df-slot 12622  df-base 12624  df-plusg 12708  df-mgm 12939  df-sgrp 12985  df-mnd 12998
This theorem is referenced by:  mndid  13006  ismndd  13018  mndpropd  13021  mhmmnd  13186
  Copyright terms: Public domain W3C validator