doc-src/AxClass/Group/Semigroups.thy
changeset 30105 37f47ea6fed1
parent 30104 b094999e1d33
parent 30101 5c6efec476ae
child 30106 351fc2f8493d
equal deleted inserted replaced
30104:b094999e1d33 30105:37f47ea6fed1
     1 
       
     2 header {* Semigroups *}
       
     3 
       
     4 theory Semigroups imports Main begin
       
     5 
       
     6 text {*
       
     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
       
     9   axioms}. Thus, type classes may be also understood as type
       
    10   predicates --- i.e.\ abstractions over a single type argument @{typ
       
    11   'a}.  Class axioms typically contain polymorphic constants that
       
    12   depend on this type @{typ 'a}.  These \emph{characteristic
       
    13   constants} behave like operations associated with the ``carrier''
       
    14   type @{typ 'a}.
       
    15 
       
    16   We illustrate these basic concepts by the following formulation of
       
    17   semigroups.
       
    18 *}
       
    19 
       
    20 consts
       
    21   times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"    (infixl "\<odot>" 70)
       
    22 axclass semigroup \<subseteq> type
       
    23   assoc: "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)"
       
    24 
       
    25 text {*
       
    26   \noindent Above we have first declared a polymorphic constant @{text
       
    27   "\<odot> \<Colon> 'a \<Rightarrow> 'a \<Rightarrow> 'a"} and then defined the class @{text semigroup} of
       
    28   all types @{text \<tau>} such that @{text "\<odot> \<Colon> \<tau> \<Rightarrow> \<tau> \<Rightarrow> \<tau>"} is indeed an
       
    29   associative operator.  The @{text assoc} axiom contains exactly one
       
    30   type variable, which is invisible in the above presentation, though.
       
    31   Also note that free term variables (like @{term x}, @{term y},
       
    32   @{term z}) are allowed for user convenience --- conceptually all of
       
    33   these are bound by outermost universal quantifiers.
       
    34 
       
    35   \medskip In general, type classes may be used to describe
       
    36   \emph{structures} with exactly one carrier @{typ 'a} and a fixed
       
    37   \emph{signature}.  Different signatures require different classes.
       
    38   Below, class @{text plus_semigroup} represents semigroups @{text
       
    39   "(\<tau>, \<oplus>\<^sup>\<tau>)"}, while the original @{text semigroup} would
       
    40   correspond to semigroups of the form @{text "(\<tau>, \<odot>\<^sup>\<tau>)"}.
       
    41 *}
       
    42 
       
    43 consts
       
    44   plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"    (infixl "\<oplus>" 70)
       
    45 axclass plus_semigroup \<subseteq> type
       
    46   assoc: "(x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
       
    47 
       
    48 text {*
       
    49   \noindent Even if classes @{text plus_semigroup} and @{text
       
    50   semigroup} both represent semigroups in a sense, they are certainly
       
    51   not quite the same.
       
    52 *}
       
    53 
       
    54 end