| author | haftmann | 
| Mon, 08 Aug 2011 22:33:36 +0200 | |
| changeset 44085 | a65e26f1427b | 
| parent 42704 | 3f19e324ff59 | 
| permissions | -rw-r--r-- | 
| 26840 | 1 | theory HOLCF_Specific | 
| 42651 | 2 | imports Base HOLCF | 
| 26840 | 3 | begin | 
| 4 | ||
| 26852 | 5 | chapter {* Isabelle/HOLCF \label{ch:holcf} *}
 | 
| 26841 | 6 | |
| 7 | section {* Mixfix syntax for continuous operations *}
 | |
| 8 | ||
| 9 | text {*
 | |
| 10 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
26852diff
changeset | 11 |     @{command_def (HOLCF) "consts"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
| 26841 | 12 |   \end{matharray}
 | 
| 13 | ||
| 14 |   HOLCF provides a separate type for continuous functions @{text "\<alpha> \<rightarrow>
 | |
| 15 |   \<beta>"}, with an explicit application operator @{term "f \<cdot> x"}.
 | |
| 16 | Isabelle mixfix syntax normally refers directly to the pure | |
| 17 |   meta-level function type @{text "\<alpha> \<Rightarrow> \<beta>"}, with application @{text "f
 | |
| 18 | x"}. | |
| 19 | ||
| 20 |   The HOLCF variant of @{command (HOLCF) "consts"} modifies that of
 | |
| 21 |   Pure Isabelle (cf.\ \secref{sec:consts}) such that declarations
 | |
| 22 | involving continuous function types are treated specifically. Any | |
| 23 | given syntax template is transformed internally, generating | |
| 24 | translation rules for the abstract and concrete representation of | |
| 25 | continuous application. Note that mixing of HOLCF and Pure | |
| 26 |   application is \emph{not} supported!
 | |
| 27 | *} | |
| 28 | ||
| 29 | ||
| 30 | section {* Recursive domains *}
 | |
| 31 | ||
| 32 | text {*
 | |
| 33 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
26852diff
changeset | 34 |     @{command_def (HOLCF) "domain"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
| 26841 | 35 |   \end{matharray}
 | 
| 36 | ||
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
30168diff
changeset | 37 |   @{rail "
 | 
| 42704 | 38 |     @@{command (HOLCF) domain} @{syntax parname}? (spec + @'and')
 | 
| 26841 | 39 | ; | 
| 40 | ||
| 42704 | 41 |     spec: @{syntax typespec} '=' (cons + '|')
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
30168diff
changeset | 42 | ; | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
30168diff
changeset | 43 |     cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}?
 | 
| 26841 | 44 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
30168diff
changeset | 45 |     dtrules: @'distinct' @{syntax thmrefs} @'inject' @{syntax thmrefs}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
30168diff
changeset | 46 |       @'induction' @{syntax thmrefs}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
30168diff
changeset | 47 | "} | 
| 26841 | 48 | |
| 49 | Recursive domains in HOLCF are analogous to datatypes in classical | |
| 50 |   HOL (cf.\ \secref{sec:hol-datatype}).  Mutual recursion is
 | |
| 51 | supported, but no nesting nor arbitrary branching. Domain | |
| 52 | constructors may be strict (default) or lazy, the latter admits to | |
| 53 | introduce infinitary objects in the typical LCF manner (e.g.\ lazy | |
| 54 |   lists).  See also \cite{MuellerNvOS99} for a general discussion of
 | |
| 55 | HOLCF domains. | |
| 56 | *} | |
| 57 | ||
| 26840 | 58 | end |