doc-src/AxClass/Group/Semigroups.thy
changeset 16417 9bc16273c2d4
parent 12338 de0f4a63baa5
equal deleted inserted replaced
16416:6061ae1f90f2 16417:9bc16273c2d4
     1 
     1 
     2 header {* Semigroups *}
     2 header {* Semigroups *}
     3 
     3 
     4 theory Semigroups = Main:
     4 theory Semigroups imports Main begin
     5 
     5 
     6 text {*
     6 text {*
     7   \medskip\noindent An axiomatic type class is simply a class of types
     7   \medskip\noindent An axiomatic type class is simply a class of types
     8   that all meet certain properties, which are also called \emph{class
     8   that all meet certain properties, which are also called \emph{class
     9   axioms}. Thus, type classes may be also understood as type
     9   axioms}. Thus, type classes may be also understood as type