| author | ballarin | 
| Mon, 08 Dec 2008 14:18:29 +0100 | |
| changeset 29019 | 8e7d6f959bd7 | 
| parent 16417 | 9bc16273c2d4 | 
| permissions | -rw-r--r-- | 
| 8906 | 1 | |
| 9146 | 2 | header {* Semigroups *}
 | 
| 8906 | 3 | |
| 16417 | 4 | theory Semigroups imports Main begin | 
| 8890 | 5 | |
| 8906 | 6 | text {*
 | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 7 | \medskip\noindent An axiomatic type class is simply a class of types | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 8 |   that all meet certain properties, which are also called \emph{class
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 9 | axioms}. Thus, type classes may be also understood as type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 10 |   predicates --- i.e.\ abstractions over a single type argument @{typ
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 11 | 'a}. Class axioms typically contain polymorphic constants that | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 12 |   depend on this type @{typ 'a}.  These \emph{characteristic
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 13 | constants} behave like operations associated with the ``carrier'' | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 14 |   type @{typ 'a}.
 | 
| 8906 | 15 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 16 | We illustrate these basic concepts by the following formulation of | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 17 | semigroups. | 
| 9146 | 18 | *} | 
| 8890 | 19 | |
| 20 | consts | |
| 10140 | 21 | times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<odot>" 70) | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 22 | axclass semigroup \<subseteq> type | 
| 10140 | 23 | assoc: "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)" | 
| 8906 | 24 | |
| 25 | text {*
 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 26 |   \noindent Above we have first declared a polymorphic constant @{text
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 27 |   "\<odot> \<Colon> 'a \<Rightarrow> 'a \<Rightarrow> 'a"} and then defined the class @{text semigroup} of
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 28 |   all types @{text \<tau>} such that @{text "\<odot> \<Colon> \<tau> \<Rightarrow> \<tau> \<Rightarrow> \<tau>"} is indeed an
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 29 |   associative operator.  The @{text assoc} axiom contains exactly one
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 30 | type variable, which is invisible in the above presentation, though. | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 31 |   Also note that free term variables (like @{term x}, @{term y},
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 32 |   @{term z}) are allowed for user convenience --- conceptually all of
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 33 | these are bound by outermost universal quantifiers. | 
| 8906 | 34 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 35 | \medskip In general, type classes may be used to describe | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 36 |   \emph{structures} with exactly one carrier @{typ 'a} and a fixed
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 37 |   \emph{signature}.  Different signatures require different classes.
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 38 |   Below, class @{text plus_semigroup} represents semigroups @{text
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 39 |   "(\<tau>, \<oplus>\<^sup>\<tau>)"}, while the original @{text semigroup} would
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 40 |   correspond to semigroups of the form @{text "(\<tau>, \<odot>\<^sup>\<tau>)"}.
 | 
| 9146 | 41 | *} | 
| 8906 | 42 | |
| 43 | consts | |
| 10140 | 44 | plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<oplus>" 70) | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 45 | axclass plus_semigroup \<subseteq> type | 
| 10140 | 46 | assoc: "(x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 8890 | 47 | |
| 8906 | 48 | text {*
 | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 49 |   \noindent Even if classes @{text plus_semigroup} and @{text
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 50 | semigroup} both represent semigroups in a sense, they are certainly | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 51 | not quite the same. | 
| 9146 | 52 | *} | 
| 8890 | 53 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 54 | end |