doc-src/IsarRef/Thy/HOLCF_Specific.thy
author blanchet
Mon, 19 Apr 2010 18:14:45 +0200
changeset 36230 43d10a494c91
parent 30168 9a20be5be90b
child 42596 6c621a9d612a
permissions -rw-r--r--
added warning about inconsistent context to Metis; it makes more sense here than in Sledgehammer, because Sledgehammer is unsound and there's no point in having people panicking about the consistency of their context when their context is in fact consistent
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26840
ec46381f149d added logic-specific sessions;
wenzelm
parents:
diff changeset
     1
theory HOLCF_Specific
ec46381f149d added logic-specific sessions;
wenzelm
parents:
diff changeset
     2
imports HOLCF
ec46381f149d added logic-specific sessions;
wenzelm
parents:
diff changeset
     3
begin
ec46381f149d added logic-specific sessions;
wenzelm
parents:
diff changeset
     4
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26841
diff changeset
     5
chapter {* Isabelle/HOLCF \label{ch:holcf} *}
26841
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
     6
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
     7
section {* Mixfix syntax for continuous operations *}
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
     8
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
     9
text {*
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    10
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 26852
diff changeset
    11
    @{command_def (HOLCF) "consts"} & : & @{text "theory \<rightarrow> theory"} \\
26841
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    12
  \end{matharray}
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    13
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    14
  HOLCF provides a separate type for continuous functions @{text "\<alpha> \<rightarrow>
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    15
  \<beta>"}, with an explicit application operator @{term "f \<cdot> x"}.
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    16
  Isabelle mixfix syntax normally refers directly to the pure
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    17
  meta-level function type @{text "\<alpha> \<Rightarrow> \<beta>"}, with application @{text "f
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    18
  x"}.
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    19
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    20
  The HOLCF variant of @{command (HOLCF) "consts"} modifies that of
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    21
  Pure Isabelle (cf.\ \secref{sec:consts}) such that declarations
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    22
  involving continuous function types are treated specifically.  Any
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    23
  given syntax template is transformed internally, generating
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    24
  translation rules for the abstract and concrete representation of
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    25
  continuous application.  Note that mixing of HOLCF and Pure
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    26
  application is \emph{not} supported!
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    27
*}
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    28
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    29
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    30
section {* Recursive domains *}
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    31
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    32
text {*
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    33
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 26852
diff changeset
    34
    @{command_def (HOLCF) "domain"} & : & @{text "theory \<rightarrow> theory"} \\
26841
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    35
  \end{matharray}
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    36
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    37
  \begin{rail}
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    38
    'domain' parname? (dmspec + 'and')
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    39
    ;
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    40
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    41
    dmspec: typespec '=' (cons + '|')
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    42
    ;
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    43
    cons: name (type *) mixfix?
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    44
    ;
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    45
    dtrules: 'distinct' thmrefs 'inject' thmrefs 'induction' thmrefs
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    46
  \end{rail}
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    47
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    48
  Recursive domains in HOLCF are analogous to datatypes in classical
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    49
  HOL (cf.\ \secref{sec:hol-datatype}).  Mutual recursion is
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    50
  supported, but no nesting nor arbitrary branching.  Domain
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    51
  constructors may be strict (default) or lazy, the latter admits to
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    52
  introduce infinitary objects in the typical LCF manner (e.g.\ lazy
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    53
  lists).  See also \cite{MuellerNvOS99} for a general discussion of
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    54
  HOLCF domains.
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    55
*}
6ac51a2f48e1 converted HOLCF specific elements;
wenzelm
parents: 26840
diff changeset
    56
26840
ec46381f149d added logic-specific sessions;
wenzelm
parents:
diff changeset
    57
end