| author | haftmann | 
| Mon, 01 Dec 2008 19:41:16 +0100 | |
| changeset 28941 | 128459bd72d2 | 
| parent 26871 | 996add9defab | 
| permissions | -rw-r--r-- | 
| 8890 | 1 | |
| 9146 | 2 | header {* Syntactic classes *}
 | 
| 8890 | 3 | |
| 16417 | 4 | theory Product imports Main begin | 
| 8890 | 5 | |
| 8903 | 6 | text {*
 | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 7 | \medskip\noindent There is still a feature of Isabelle's type system | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 8 | left that we have not yet discussed. When declaring polymorphic | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 9 |   constants @{text "c \<Colon> \<sigma>"}, the type variables occurring in @{text \<sigma>}
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 10 | may be constrained by type classes (or even general sorts) in an | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 11 | arbitrary way. Note that by default, in Isabelle/HOL the | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 12 |   declaration @{text "\<odot> \<Colon> 'a \<Rightarrow> 'a \<Rightarrow> 'a"} is actually an abbreviation
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 13 |   for @{text "\<odot> \<Colon> 'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> 'a"} Since class @{text type} is the
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 14 | universal class of HOL, this is not really a constraint at all. | 
| 8903 | 15 | |
| 10140 | 16 |  The @{text product} class below provides a less degenerate example of
 | 
| 8903 | 17 | syntactic type classes. | 
| 9146 | 18 | *} | 
| 8903 | 19 | |
| 8890 | 20 | axclass | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 21 | product \<subseteq> type | 
| 8890 | 22 | consts | 
| 10140 | 23 | product :: "'a\<Colon>product \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<odot>" 70) | 
| 8890 | 24 | |
| 8903 | 25 | text {*
 | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 26 |   Here class @{text product} is defined as subclass of @{text type}
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 27 | without any additional axioms. This effects in logical equivalence | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 28 |   of @{text product} and @{text type}, as is reflected by the trivial
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 29 | introduction rule generated for this definition. | 
| 8903 | 30 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 31 |   \medskip So what is the difference of declaring @{text "\<odot> \<Colon>
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 32 |   'a\<Colon>product \<Rightarrow> 'a \<Rightarrow> 'a"} vs.\ declaring @{text "\<odot> \<Colon> 'a\<Colon>type \<Rightarrow> 'a \<Rightarrow>
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 33 |   'a"} anyway?  In this particular case where @{text "product \<equiv>
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 34 | type"}, it should be obvious that both declarations are the same | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 35 | from the logic's point of view. It even makes the most sense to | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 36 | remove sort constraints from constant declarations, as far as the | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 37 |   purely logical meaning is concerned \cite{Wenzel:1997:TPHOL}.
 | 
| 8903 | 38 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 39 | On the other hand there are syntactic differences, of course. | 
| 26871 | 40 |  Constants @{text \<odot>} on some type @{text \<tau>} are rejected by the
 | 
| 41 |  type-checker, unless the arity @{text "\<tau> \<Colon> product"} is part of the
 | |
| 42 | type signature. In our example, this arity may be always added when | |
| 43 |  required by means of an \isakeyword{instance} with the default proof
 | |
| 44 | (double-dot). | |
| 8903 | 45 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 46 | \medskip Thus, we may observe the following discipline of using | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 47 | syntactic classes. Overloaded polymorphic constants have their type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 48 | arguments restricted to an associated (logically trivial) class | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 49 |   @{text c}.  Only immediately before \emph{specifying} these
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 50 |   constants on a certain type @{text \<tau>} do we instantiate @{text "\<tau> \<Colon>
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 51 | c"}. | 
| 8903 | 52 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 53 |   This is done for class @{text product} and type @{typ bool} as
 | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 54 | follows. | 
| 9146 | 55 | *} | 
| 8903 | 56 | |
| 10309 | 57 | instance bool :: product .. | 
| 9306 | 58 | defs (overloaded) | 
| 10140 | 59 | product_bool_def: "x \<odot> y \<equiv> x \<and> y" | 
| 8890 | 60 | |
| 8903 | 61 | text {*
 | 
| 10140 | 62 |  The definition @{text prod_bool_def} becomes syntactically
 | 
| 63 |  well-formed only after the arity @{text "bool \<Colon> product"} is made
 | |
| 64 | known to the type checker. | |
| 8903 | 65 | |
| 26871 | 66 |  \medskip It is very important to see that above \isakeyword{defs} are
 | 
| 67 |  not directly connected with \isakeyword{instance} at all!  We were
 | |
| 68 |  just following our convention to specify @{text \<odot>} on @{typ bool}
 | |
| 69 |  after having instantiated @{text "bool \<Colon> product"}.  Isabelle does
 | |
| 70 | not require these definitions, which is in contrast to programming | |
| 71 |  languages like Haskell \cite{haskell-report}.
 | |
| 8903 | 72 | |
| 73 | \medskip While Isabelle type classes and those of Haskell are almost | |
| 74 | the same as far as type-checking and type inference are concerned, | |
| 8907 | 75 | there are important semantic differences. Haskell classes require | 
| 76 |  their instances to \emph{provide operations} of certain \emph{names}.
 | |
| 8903 | 77 |  Therefore, its \texttt{instance} has a \texttt{where} part that tells
 | 
| 78 | the system what these ``member functions'' should be. | |
| 79 | ||
| 10140 | 80 |  This style of \texttt{instance} would not make much sense in
 | 
| 81 | Isabelle's meta-logic, because there is no internal notion of | |
| 82 | ``providing operations'' or even ``names of functions''. | |
| 9146 | 83 | *} | 
| 8903 | 84 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11099diff
changeset | 85 | end |