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