src/Doc/Tutorial/Types/Overloading.thy
author wenzelm
Sat, 05 Apr 2014 15:03:40 +0200
changeset 56421 1ffd7eaa778b
parent 54665 617ddc60f914
child 58842 22b87ab47d3b
permissions -rw-r--r--
updated to jedit_build-20140405: Code2HTML.jar, CommonControls.jar, Console.jar, kappalayout.jar, Navigator.jar, SideKick.jar, doc with jEdit manuals (ant dist-manuals);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31680
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
     1
(*<*)theory Overloading imports Main Setup begin
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
     2
36176
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 31707
diff changeset
     3
hide_class (open) plus (*>*)
31680
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
     4
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
     5
text {* Type classes allow \emph{overloading}; thus a constant may
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
     6
have multiple definitions at non-overlapping types. *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
     7
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
     8
subsubsection {* Overloading *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
     9
54665
wenzelm
parents: 48985
diff changeset
    10
text {* We can introduce a binary infix addition operator @{text "\<oplus>"}
31680
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    11
for arbitrary types by means of a type class: *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    12
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    13
class plus =
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    14
  fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<oplus>" 70)
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    15
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    16
text {* \noindent This introduces a new class @{class [source] plus},
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    17
along with a constant @{const [source] plus} with nice infix syntax.
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    18
@{const [source] plus} is also named \emph{class operation}.  The type
31682
358cdcdf56d2 corrected some issues
haftmann
parents: 31680
diff changeset
    19
of @{const [source] plus} carries a class constraint @{typ [source] "'a
31680
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    20
:: plus"} on its type variable, meaning that only types of class
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    21
@{class [source] plus} can be instantiated for @{typ [source] "'a"}.
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    22
To breathe life into @{class [source] plus} we need to declare a type
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    23
to be an \bfindex{instance} of @{class [source] plus}: *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    24
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    25
instantiation nat :: plus
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    26
begin
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    27
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    28
text {* \noindent Command \isacommand{instantiation} opens a local
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    29
theory context.  Here we can now instantiate @{const [source] plus} on
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    30
@{typ nat}: *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    31
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    32
primrec plus_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    33
    "(0::nat) \<oplus> n = n"
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    34
  | "Suc m \<oplus> n = Suc (m \<oplus> n)"
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    35
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    36
text {* \noindent Note that the name @{const [source] plus} carries a
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    37
suffix @{text "_nat"}; by default, the local name of a class operation
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    38
@{text f} to be instantiated on type constructor @{text \<kappa>} is mangled
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    39
as @{text f_\<kappa>}.  In case of uncertainty, these names may be inspected
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    40
using the @{command "print_context"} command or the corresponding
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    41
ProofGeneral button.
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    42
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    43
Although class @{class [source] plus} has no axioms, the instantiation must be
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    44
formally concluded by a (trivial) instantiation proof ``..'': *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    45
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    46
instance ..
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    47
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    48
text {* \noindent More interesting \isacommand{instance} proofs will
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    49
arise below.
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    50
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    51
The instantiation is finished by an explicit *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    52
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    53
end
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    54
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    55
text {* \noindent From now on, terms like @{term "Suc (m \<oplus> 2)"} are
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    56
legal. *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    57
38325
6daf896bca5e * -> prod
haftmann
parents: 36176
diff changeset
    58
instantiation prod :: (plus, plus) plus
31680
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    59
begin
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    60
38325
6daf896bca5e * -> prod
haftmann
parents: 36176
diff changeset
    61
text {* \noindent Here we instantiate the product type @{type prod} to
31680
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    62
class @{class [source] plus}, given that its type arguments are of
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    63
class @{class [source] plus}: *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    64
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    65
fun plus_prod :: "'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b" where
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    66
  "(x, y) \<oplus> (w, z) = (x \<oplus> w, y \<oplus> z)"
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    67
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    68
text {* \noindent Obviously, overloaded specifications may include
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    69
recursion over the syntactic structure of types. *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    70
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    71
instance ..
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    72
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    73
end
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    74
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    75
text {* \noindent This way we have encoded the canonical lifting of
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    76
binary operations to products by means of type classes. *}
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    77
614a8c4c9c0f reworked section on type classes
haftmann
parents:
diff changeset
    78
(*<*)end(*>*)