doc-src/IsarRef/Thy/HOL_Specific.thy
author blanchet
Mon, 06 Jun 2011 20:36:35 +0200
changeset 43187 95bd1ef1331a
parent 43040 665623e695ea
child 43270 bc72c1ccc89e
permissions -rw-r--r--
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26840
ec46381f149d added logic-specific sessions;
wenzelm
parents:
diff changeset
     1
theory HOL_Specific
42651
e3fdb7c96be5 formal Base theory;
wenzelm
parents: 42627
diff changeset
     2
imports Base Main
26840
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: 26849
diff changeset
     5
chapter {* Isabelle/HOL \label{ch:hol} *}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
     6
35744
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
     7
section {* Typedef axiomatization \label{sec:hol-typedef} *}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
     8
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
     9
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    10
  \begin{matharray}{rcl}
35744
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    11
    @{command_def (HOL) "typedef"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    12
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    13
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
    14
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
    15
    @@{command (HOL) typedef} altname? abstype '=' repset
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    16
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    17
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
    18
    altname: '(' (@{syntax name} | @'open' | @'open' @{syntax name}) ')'
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    19
    ;
42705
528a2ba8fa74 tuned some syntax names;
wenzelm
parents: 42704
diff changeset
    20
    abstype: @{syntax typespec_sorts} @{syntax mixfix}?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    21
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
    22
    repset: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
    23
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    24
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    25
  \begin{description}
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
    26
35744
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    27
  \item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"}
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    28
  axiomatizes a Gordon/HOL-style type definition in the background
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    29
  theory of the current context, depending on a non-emptiness result
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    30
  of the set @{text A} (which needs to be proven interactively).
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    31
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    32
  The raw type may not depend on parameters or assumptions of the
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    33
  context --- this is logically impossible in Isabelle/HOL --- but the
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    34
  non-emptiness property can be local, potentially resulting in
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    35
  multiple interpretations in target contexts.  Thus the established
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    36
  bijection between the representing set @{text A} and the new type
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    37
  @{text t} may semantically depend on local assumptions.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
    38
35744
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    39
  By default, @{command (HOL) "typedef"} defines both a type @{text t}
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    40
  and a set (term constant) of the same name, unless an alternative
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    41
  base name is given in parentheses, or the ``@{text "(open)"}''
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    42
  declaration is used to suppress a separate constant definition
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    43
  altogether.  The injection from type to set is called @{text Rep_t},
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    44
  its inverse @{text Abs_t} --- this may be changed via an explicit
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    45
  @{keyword (HOL) "morphisms"} declaration.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
    46
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    47
  Theorems @{text Rep_t}, @{text Rep_t_inverse}, and @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    48
  Abs_t_inverse} provide the most basic characterization as a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    49
  corresponding injection/surjection pair (in both directions).  Rules
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    50
  @{text Rep_t_inject} and @{text Abs_t_inject} provide a slightly
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    51
  more convenient view on the injectivity part, suitable for automated
26894
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26860
diff changeset
    52
  proof tools (e.g.\ in @{attribute simp} or @{attribute iff}
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26860
diff changeset
    53
  declarations).  Rules @{text Rep_t_cases}/@{text Rep_t_induct}, and
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26860
diff changeset
    54
  @{text Abs_t_cases}/@{text Abs_t_induct} provide alternative views
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26860
diff changeset
    55
  on surjectivity; these are already declared as set or type rules for
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    56
  the generic @{method cases} and @{method induct} methods.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
    57
35744
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    58
  An alternative name for the set definition (and other derived
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    59
  entities) may be specified in parentheses; the default is to use
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    60
  @{text t} as indicated before.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    61
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    62
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    63
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    64
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    65
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    66
section {* Adhoc tuples *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    67
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    68
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    69
  \begin{matharray}{rcl}
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
    70
    @{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    71
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    72
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
    73
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
    74
    @@{attribute (HOL) split_format} ('(' 'complete' ')')?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
    75
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    76
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    77
  \begin{description}
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
    78
40388
cb9fd7dd641c abolished obscure goal variant of [split_format] -- unused (cf. d1c14898fd04), unrelated to '(complete)' variant, and not at all canonical
krauss
parents: 40255
diff changeset
    79
  \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    80
  arguments in function applications to be represented canonically
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    81
  according to their tuple type structure.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    82
40388
cb9fd7dd641c abolished obscure goal variant of [split_format] -- unused (cf. d1c14898fd04), unrelated to '(complete)' variant, and not at all canonical
krauss
parents: 40255
diff changeset
    83
  Note that this operation tends to invent funny names for new local
cb9fd7dd641c abolished obscure goal variant of [split_format] -- unused (cf. d1c14898fd04), unrelated to '(complete)' variant, and not at all canonical
krauss
parents: 40255
diff changeset
    84
  parameters introduced.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    85
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    86
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    87
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    88
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    89
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    90
section {* Records \label{sec:hol-record} *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    91
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    92
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    93
  In principle, records merely generalize the concept of tuples, where
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    94
  components may be addressed by labels instead of just position.  The
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    95
  logical infrastructure of records in Isabelle/HOL is slightly more
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    96
  advanced, though, supporting truly extensible record schemes.  This
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    97
  admits operations that are polymorphic with respect to record
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    98
  extension, yielding ``object-oriented'' effects like (single)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    99
  inheritance.  See also \cite{NaraschewskiW-TPHOLs98} for more
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   100
  details on object-oriented verification and record subtyping in HOL.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   101
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   102
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   103
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   104
subsection {* Basic concepts *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   105
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   106
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   107
  Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   108
  at the level of terms and types.  The notation is as follows:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   109
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   110
  \begin{center}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   111
  \begin{tabular}{l|l|l}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   112
    & record terms & record types \\ \hline
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   113
    fixed & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   114
    schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   115
      @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   116
  \end{tabular}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   117
  \end{center}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   118
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   119
  \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   120
  "(| x = a |)"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   121
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   122
  A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   123
  @{text a} and field @{text y} of value @{text b}.  The corresponding
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   124
  type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   125
  and @{text "b :: B"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   126
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   127
  A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   128
  @{text x} and @{text y} as before, but also possibly further fields
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   129
  as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   130
  of the syntax).  The improper field ``@{text "\<dots>"}'' of a record
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   131
  scheme is called the \emph{more part}.  Logically it is just a free
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   132
  variable, which is occasionally referred to as ``row variable'' in
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   133
  the literature.  The more part of a record scheme may be
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   134
  instantiated by zero or more further components.  For example, the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   135
  previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z =
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   136
  c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   137
  Fixed records are special instances of record schemes, where
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   138
  ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   139
  element.  In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   140
  for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   141
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   142
  \medskip Two key observations make extensible records in a simply
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   143
  typed language like HOL work out:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   144
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   145
  \begin{enumerate}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   146
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   147
  \item the more part is internalized, as a free term or type
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   148
  variable,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   149
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   150
  \item field names are externalized, they cannot be accessed within
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   151
  the logic as first-class values.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   152
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   153
  \end{enumerate}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   154
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   155
  \medskip In Isabelle/HOL record types have to be defined explicitly,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   156
  fixing their field names and types, and their (optional) parent
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   157
  record.  Afterwards, records may be formed using above syntax, while
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   158
  obeying the canonical order of fields as given by their declaration.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   159
  The record package provides several standard operations like
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   160
  selectors and updates.  The common setup for various generic proof
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   161
  tools enable succinct reasoning patterns.  See also the Isabelle/HOL
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   162
  tutorial \cite{isabelle-hol-book} for further instructions on using
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   163
  records in practice.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   164
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   165
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   166
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   167
subsection {* Record specifications *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   168
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   169
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   170
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   171
    @{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   172
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   173
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   174
  @{rail "
42705
528a2ba8fa74 tuned some syntax names;
wenzelm
parents: 42704
diff changeset
   175
    @@{command (HOL) record} @{syntax typespec_sorts} '=' \\
42704
3f19e324ff59 tuned rail diagrams and layout;
wenzelm
parents: 42652
diff changeset
   176
      (@{syntax type} '+')? (@{syntax constdecl} +)
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   177
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   178
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   179
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   180
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   181
  \item @{command (HOL) "record"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   182
  \<dots> c\<^sub>n :: \<sigma>\<^sub>n"} defines extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"},
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   183
  derived from the optional parent record @{text "\<tau>"} by adding new
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   184
  field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   185
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   186
  The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   187
  covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   188
  \<alpha>\<^sub>m"}.  Type constructor @{text t} has to be new, while @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   189
  \<tau>} needs to specify an instance of an existing record type.  At
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   190
  least one new field @{text "c\<^sub>i"} has to be specified.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   191
  Basically, field names need to belong to a unique record.  This is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   192
  not a real restriction in practice, since fields are qualified by
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   193
  the record name internally.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   194
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   195
  The parent record specification @{text \<tau>} is optional; if omitted
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   196
  @{text t} becomes a root record.  The hierarchy of all records
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   197
  declared within a theory context forms a forest structure, i.e.\ a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   198
  set of trees starting with a root record each.  There is no way to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   199
  merge multiple parent records!
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   200
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   201
  For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   202
  type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   203
  \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   204
  "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   205
  @{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> ::
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   206
  \<zeta>\<rparr>"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   207
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   208
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   209
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   210
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   211
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   212
subsection {* Record operations *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   213
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   214
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   215
  Any record definition of the form presented above produces certain
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   216
  standard operations.  Selectors and updates are provided for any
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   217
  field, including the improper one ``@{text more}''.  There are also
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   218
  cumulative record constructor functions.  To simplify the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   219
  presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   220
  \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   221
  \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   222
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   223
  \medskip \textbf{Selectors} and \textbf{updates} are available for
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   224
  any field (including ``@{text more}''):
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   225
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   226
  \begin{matharray}{lll}
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   227
    @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   228
    @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   229
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   230
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   231
  There is special syntax for application of updates: @{text "r\<lparr>x :=
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   232
  a\<rparr>"} abbreviates term @{text "x_update a r"}.  Further notation for
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   233
  repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z :=
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   234
  c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}.  Note that
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   235
  because of postfix notation the order of fields shown here is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   236
  reverse than in the actual term.  Since repeated updates are just
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   237
  function applications, fields may be freely permuted in @{text "\<lparr>x
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   238
  := a, y := b, z := c\<rparr>"}, as far as logical equality is concerned.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   239
  Thus commutativity of independent updates can be proven within the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   240
  logic for any two fields, but not as a general theorem.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   241
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   242
  \medskip The \textbf{make} operation provides a cumulative record
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   243
  constructor function:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   244
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   245
  \begin{matharray}{lll}
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   246
    @{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   247
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   248
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   249
  \medskip We now reconsider the case of non-root records, which are
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   250
  derived of some parent.  In general, the latter may depend on
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   251
  another parent as well, resulting in a list of \emph{ancestor
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   252
  records}.  Appending the lists of fields of all ancestors results in
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   253
  a certain field prefix.  The record package automatically takes care
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   254
  of this by lifting operations over this context of ancestor fields.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   255
  Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   256
  fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"},
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   257
  the above record operations will get the following types:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   258
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   259
  \medskip
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   260
  \begin{tabular}{lll}
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   261
    @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   262
    @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow>
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   263
      \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   264
      \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   265
    @{text "t.make"} & @{text "::"} & @{text "\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow>
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   266
      \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   267
  \end{tabular}
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   268
  \medskip
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   269
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   270
  \noindent Some further operations address the extension aspect of a
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   271
  derived record scheme specifically: @{text "t.fields"} produces a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   272
  record fragment consisting of exactly the new fields introduced here
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   273
  (the result may serve as a more part elsewhere); @{text "t.extend"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   274
  takes a fixed record and adds a given more part; @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   275
  "t.truncate"} restricts a record scheme to a fixed record.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   276
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   277
  \medskip
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   278
  \begin{tabular}{lll}
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   279
    @{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   280
    @{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   281
      \<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   282
    @{text "t.truncate"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   283
  \end{tabular}
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   284
  \medskip
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   285
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   286
  \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   287
  for root records.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   288
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   289
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   290
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   291
subsection {* Derived rules and proof tools *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   292
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   293
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   294
  The record package proves several results internally, declaring
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   295
  these facts to appropriate proof tools.  This enables users to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   296
  reason about record structures quite conveniently.  Assume that
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   297
  @{text t} is a record type as specified above.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   298
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   299
  \begin{enumerate}
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   300
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   301
  \item Standard conversions for selectors or updates applied to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   302
  record constructor terms are made part of the default Simplifier
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   303
  context; thus proofs by reduction of basic operations merely require
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   304
  the @{method simp} method without further arguments.  These rules
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   305
  are available as @{text "t.simps"}, too.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   306
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   307
  \item Selectors applied to updated records are automatically reduced
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   308
  by an internal simplification procedure, which is also part of the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   309
  standard Simplifier setup.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   310
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   311
  \item Inject equations of a form analogous to @{prop "(x, y) = (x',
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   312
  y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   313
  Reasoner as @{attribute iff} rules.  These rules are available as
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   314
  @{text "t.iffs"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   315
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   316
  \item The introduction rule for record equality analogous to @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   317
  "x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   318
  and as the basic rule context as ``@{attribute intro}@{text "?"}''.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   319
  The rule is called @{text "t.equality"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   320
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   321
  \item Representations of arbitrary record expressions as canonical
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   322
  constructor terms are provided both in @{method cases} and @{method
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   323
  induct} format (cf.\ the generic proof methods of the same name,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   324
  \secref{sec:cases-induct}).  Several variations are available, for
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   325
  fixed records, record schemes, more parts etc.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   326
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   327
  The generic proof methods are sufficiently smart to pick the most
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   328
  sensible rule according to the type of the indicated record
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   329
  expression: users just need to apply something like ``@{text "(cases
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   330
  r)"}'' to a certain proof problem.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   331
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   332
  \item The derived record operations @{text "t.make"}, @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   333
  "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   334
  treated automatically, but usually need to be expanded by hand,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   335
  using the collective fact @{text "t.defs"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   336
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   337
  \end{enumerate}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   338
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   339
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   340
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   341
section {* Datatypes \label{sec:hol-datatype} *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   342
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   343
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   344
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   345
    @{command_def (HOL) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\
41396
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   346
    @{command_def (HOL) "rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   347
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   348
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   349
  @{rail "
42704
3f19e324ff59 tuned rail diagrams and layout;
wenzelm
parents: 42652
diff changeset
   350
    @@{command (HOL) datatype} (spec + @'and')
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   351
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   352
    @@{command (HOL) rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +)
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   353
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   354
42704
3f19e324ff59 tuned rail diagrams and layout;
wenzelm
parents: 42652
diff changeset
   355
    spec: @{syntax parname}? @{syntax typespec} @{syntax mixfix}? '=' (cons + '|')
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   356
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   357
    cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   358
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   359
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   360
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   361
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   362
  \item @{command (HOL) "datatype"} defines inductive datatypes in
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   363
  HOL.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   364
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   365
  \item @{command (HOL) "rep_datatype"} represents existing types as
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   366
  inductive ones, generating the standard infrastructure of derived
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   367
  concepts (primitive recursion etc.).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   368
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   369
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   370
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   371
  The induction and exhaustion theorems generated provide case names
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   372
  according to the constructors involved, while parameters are named
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   373
  after the types (see also \secref{sec:cases-induct}).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   374
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   375
  See \cite{isabelle-HOL} for more details on datatypes, but beware of
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   376
  the old-style theory syntax being used there!  Apart from proper
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   377
  proof methods for case-analysis and induction, there are also
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   378
  emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   379
  induct_tac} available, see \secref{sec:hol-induct-tac}; these admit
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   380
  to refer directly to the internal structure of subgoals (including
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   381
  internally bound parameters).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   382
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   383
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   384
41396
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   385
section {* Functorial structure of types *}
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   386
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   387
text {*
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   388
  \begin{matharray}{rcl}
41505
6d19301074cf "enriched_type" replaces less specific "type_lifting"
haftmann
parents: 41396
diff changeset
   389
    @{command_def (HOL) "enriched_type"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
41396
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   390
  \end{matharray}
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   391
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   392
  @{rail "
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   393
    @@{command (HOL) enriched_type} (@{syntax name} ':')? @{syntax term}
41396
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   394
    ;
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   395
  "}
41396
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   396
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   397
  \begin{description}
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   398
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   399
  \item @{command (HOL) "enriched_type"}~@{text "prefix: m"} allows to
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   400
  prove and register properties about the functorial structure of type
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   401
  constructors.  These properties then can be used by other packages
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   402
  to deal with those type constructors in certain type constructions.
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   403
  Characteristic theorems are noted in the current local theory.  By
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   404
  default, they are prefixed with the base name of the type
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   405
  constructor, an explicit prefix can be given alternatively.
41396
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   406
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   407
  The given term @{text "m"} is considered as \emph{mapper} for the
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   408
  corresponding type constructor and must conform to the following
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   409
  type pattern:
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   410
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   411
  \begin{matharray}{lll}
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   412
    @{text "m"} & @{text "::"} &
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   413
      @{text "\<sigma>\<^isub>1 \<Rightarrow> \<dots> \<sigma>\<^isub>k \<Rightarrow> (\<^vec>\<alpha>\<^isub>n) t \<Rightarrow> (\<^vec>\<beta>\<^isub>n) t"} \\
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   414
  \end{matharray}
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   415
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   416
  \noindent where @{text t} is the type constructor, @{text
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   417
  "\<^vec>\<alpha>\<^isub>n"} and @{text "\<^vec>\<beta>\<^isub>n"} are distinct
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   418
  type variables free in the local theory and @{text "\<sigma>\<^isub>1"},
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   419
  \ldots, @{text "\<sigma>\<^isub>k"} is a subsequence of @{text "\<alpha>\<^isub>1 \<Rightarrow>
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   420
  \<beta>\<^isub>1"}, @{text "\<beta>\<^isub>1 \<Rightarrow> \<alpha>\<^isub>1"}, \ldots,
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   421
  @{text "\<alpha>\<^isub>n \<Rightarrow> \<beta>\<^isub>n"}, @{text "\<beta>\<^isub>n \<Rightarrow>
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   422
  \<alpha>\<^isub>n"}.
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   423
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   424
  \end{description}
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   425
*}
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   426
5379e4a85a66 documentation stub on type_lifting
haftmann
parents: 40918
diff changeset
   427
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   428
section {* Recursive functions \label{sec:recursion} *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   429
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   430
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   431
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   432
    @{command_def (HOL) "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   433
    @{command_def (HOL) "fun"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   434
    @{command_def (HOL) "function"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   435
    @{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   436
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   437
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   438
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   439
    @@{command (HOL) primrec} @{syntax target}? @{syntax \"fixes\"} @'where' equations
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   440
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   441
    (@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   442
      @{syntax \"fixes\"} \\ @'where' equations
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   443
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   444
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   445
    equations: (@{syntax thmdecl}? @{syntax prop} + '|')
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   446
    ;
41846
b368a7aee46a removed support for tail-recursion from function package (now implemented by partial_function)
krauss
parents: 41506
diff changeset
   447
    functionopts: '(' (('sequential' | 'domintros') + ',') ')'
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   448
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   449
    @@{command (HOL) termination} @{syntax term}?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   450
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   451
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   452
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   453
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   454
  \item @{command (HOL) "primrec"} defines primitive recursive
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   455
  functions over datatypes, see also \cite{isabelle-HOL}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   456
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   457
  \item @{command (HOL) "function"} defines functions by general
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   458
  wellfounded recursion. A detailed description with examples can be
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   459
  found in \cite{isabelle-function}. The function is specified by a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   460
  set of (possibly conditional) recursive equations with arbitrary
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   461
  pattern matching. The command generates proof obligations for the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   462
  completeness and the compatibility of patterns.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   463
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   464
  The defined function is considered partial, and the resulting
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   465
  simplification rules (named @{text "f.psimps"}) and induction rule
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   466
  (named @{text "f.pinduct"}) are guarded by a generated domain
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   467
  predicate @{text "f_dom"}. The @{command (HOL) "termination"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   468
  command can then be used to establish that the function is total.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   469
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   470
  \item @{command (HOL) "fun"} is a shorthand notation for ``@{command
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   471
  (HOL) "function"}~@{text "(sequential)"}, followed by automated
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   472
  proof attempts regarding pattern matching and termination.  See
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   473
  \cite{isabelle-function} for further details.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   474
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   475
  \item @{command (HOL) "termination"}~@{text f} commences a
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   476
  termination proof for the previously defined function @{text f}.  If
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   477
  this is omitted, the command refers to the most recent function
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   478
  definition.  After the proof is closed, the recursive equations and
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   479
  the induction principle is established.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   480
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   481
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   482
27452
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
   483
  Recursive definitions introduced by the @{command (HOL) "function"}
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
   484
  command accommodate
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   485
  reasoning by induction (cf.\ \secref{sec:cases-induct}): rule @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   486
  "c.induct"} (where @{text c} is the name of the function definition)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   487
  refers to a specific induction rule, with parameters named according
33857
0cb5002c52db clarified; checked
krauss
parents: 31998
diff changeset
   488
  to the user-specified equations. Cases are numbered (starting from 1).
0cb5002c52db clarified; checked
krauss
parents: 31998
diff changeset
   489
0cb5002c52db clarified; checked
krauss
parents: 31998
diff changeset
   490
  For @{command (HOL) "primrec"}, the induction principle coincides
27452
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
   491
  with structural recursion on the datatype the recursion is carried
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
   492
  out.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   493
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   494
  The equations provided by these packages may be referred later as
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   495
  theorem list @{text "f.simps"}, where @{text f} is the (collective)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   496
  name of the functions defined.  Individual equations may be named
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   497
  explicitly as well.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   498
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   499
  The @{command (HOL) "function"} command accepts the following
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   500
  options.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   501
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   502
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   503
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   504
  \item @{text sequential} enables a preprocessor which disambiguates
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   505
  overlapping patterns by making them mutually disjoint.  Earlier
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   506
  equations take precedence over later ones.  This allows to give the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   507
  specification in a format very similar to functional programming.
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   508
  Note that the resulting simplification and induction rules
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   509
  correspond to the transformed specification, not the one given
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   510
  originally. This usually means that each equation given by the user
36139
0c2538afe8e8 Spelling error: theroems -> theorems
hoelzl
parents: 35841
diff changeset
   511
  may result in several theorems.  Also note that this automatic
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   512
  transformation only works for ML-style datatype patterns.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   513
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   514
  \item @{text domintros} enables the automated generation of
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   515
  introduction rules for the domain predicate. While mostly not
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   516
  needed, they can be helpful in some proofs about partial functions.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   517
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   518
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   519
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   520
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   521
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   522
subsection {* Proof methods related to recursive definitions *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   523
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   524
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   525
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   526
    @{method_def (HOL) pat_completeness} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   527
    @{method_def (HOL) relation} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   528
    @{method_def (HOL) lexicographic_order} & : & @{text method} \\
33858
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   529
    @{method_def (HOL) size_change} & : & @{text method} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   530
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   531
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   532
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   533
    @@{method (HOL) relation} @{syntax term}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   534
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   535
    @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * )
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   536
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   537
    @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) )
33858
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   538
    ;
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   539
    orders: ( 'max' | 'min' | 'ms' ) *
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   540
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   541
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   542
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   543
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   544
  \item @{method (HOL) pat_completeness} is a specialized method to
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   545
  solve goals regarding the completeness of pattern matching, as
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   546
  required by the @{command (HOL) "function"} package (cf.\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   547
  \cite{isabelle-function}).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   548
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   549
  \item @{method (HOL) relation}~@{text R} introduces a termination
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   550
  proof using the relation @{text R}.  The resulting proof state will
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   551
  contain goals expressing that @{text R} is wellfounded, and that the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   552
  arguments of recursive calls decrease with respect to @{text R}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   553
  Usually, this method is used as the initial proof step of manual
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   554
  termination proofs.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   555
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   556
  \item @{method (HOL) "lexicographic_order"} attempts a fully
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   557
  automated termination proof by searching for a lexicographic
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   558
  combination of size measures on the arguments of the function. The
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   559
  method accepts the same arguments as the @{method auto} method,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   560
  which it uses internally to prove local descents.  The same context
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   561
  modifiers as for @{method auto} are accepted, see
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   562
  \secref{sec:clasimp}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   563
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   564
  In case of failure, extensive information is printed, which can help
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   565
  to analyse the situation (cf.\ \cite{isabelle-function}).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   566
33858
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   567
  \item @{method (HOL) "size_change"} also works on termination goals,
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   568
  using a variation of the size-change principle, together with a
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   569
  graph decomposition technique (see \cite{krauss_phd} for details).
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   570
  Three kinds of orders are used internally: @{text max}, @{text min},
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   571
  and @{text ms} (multiset), which is only available when the theory
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   572
  @{text Multiset} is loaded. When no order kinds are given, they are
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   573
  tried in order. The search for a termination proof uses SAT solving
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   574
  internally.
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   575
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   576
 For local descent proofs, the same context modifiers as for @{method
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   577
  auto} are accepted, see \secref{sec:clasimp}.
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   578
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   579
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   580
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   581
40171
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   582
subsection {* Functions with explicit partiality *}
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   583
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   584
text {*
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   585
  \begin{matharray}{rcl}
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   586
    @{command_def (HOL) "partial_function"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   587
    @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   588
  \end{matharray}
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   589
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   590
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   591
    @@{command (HOL) partial_function} @{syntax target}?
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   592
      '(' @{syntax nameref} ')' @{syntax \"fixes\"} \\
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   593
      @'where' @{syntax thmdecl}? @{syntax prop}
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   594
  "}
40171
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   595
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   596
  \begin{description}
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   597
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   598
  \item @{command (HOL) "partial_function"}~@{text "(mode)"} defines
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   599
  recursive functions based on fixpoints in complete partial
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   600
  orders. No termination proof is required from the user or
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   601
  constructed internally. Instead, the possibility of non-termination
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   602
  is modelled explicitly in the result type, which contains an
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   603
  explicit bottom element.
40171
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   604
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   605
  Pattern matching and mutual recursion are currently not supported.
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   606
  Thus, the specification consists of a single function described by a
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   607
  single recursive equation.
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   608
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   609
  There are no fixed syntactic restrictions on the body of the
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   610
  function, but the induced functional must be provably monotonic
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   611
  wrt.\ the underlying order.  The monotonicitity proof is performed
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   612
  internally, and the definition is rejected when it fails. The proof
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   613
  can be influenced by declaring hints using the
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   614
  @{attribute (HOL) partial_function_mono} attribute.
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   615
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   616
  The mandatory @{text mode} argument specifies the mode of operation
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   617
  of the command, which directly corresponds to a complete partial
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   618
  order on the result type. By default, the following modes are
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   619
  defined:
40171
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   620
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   621
  \begin{description}
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   622
  \item @{text option} defines functions that map into the @{type
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   623
  option} type. Here, the value @{term None} is used to model a
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   624
  non-terminating computation. Monotonicity requires that if @{term
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   625
  None} is returned by a recursive call, then the overall result
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   626
  must also be @{term None}. This is best achieved through the use of
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   627
  the monadic operator @{const "Option.bind"}.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   628
40171
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   629
  \item @{text tailrec} defines functions with an arbitrary result
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   630
  type and uses the slightly degenerated partial order where @{term
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   631
  "undefined"} is the bottom element.  Now, monotonicity requires that
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   632
  if @{term undefined} is returned by a recursive call, then the
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   633
  overall result must also be @{term undefined}. In practice, this is
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   634
  only satisfied when each recursive call is a tail call, whose result
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   635
  is directly returned. Thus, this mode of operation allows the
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   636
  definition of arbitrary tail-recursive functions.
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   637
  \end{description}
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   638
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   639
  Experienced users may define new modes by instantiating the locale
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   640
  @{const "partial_function_definitions"} appropriately.
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   641
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   642
  \item @{attribute (HOL) partial_function_mono} declares rules for
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   643
  use in the internal monononicity proofs of partial function
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   644
  definitions.
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   645
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   646
  \end{description}
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   647
1fa547166a1d basic documentation for command partial_function
krauss
parents: 40170
diff changeset
   648
*}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   649
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   650
subsection {* Old-style recursive function definitions (TFL) *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   651
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   652
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   653
  The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   654
  "recdef_tc"} for defining recursive are mostly obsolete; @{command
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   655
  (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   656
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   657
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   658
    @{command_def (HOL) "recdef"} & : & @{text "theory \<rightarrow> theory)"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   659
    @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   660
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   661
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   662
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   663
    @@{command (HOL) recdef} ('(' @'permissive' ')')? \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   664
      @{syntax name} @{syntax term} (@{syntax prop} +) hints?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   665
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   666
    recdeftc @{syntax thmdecl}? tc
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   667
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   668
    hints: '(' @'hints' ( recdefmod * ) ')'
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   669
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   670
    recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf')
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   671
      (() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   672
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   673
    tc: @{syntax nameref} ('(' @{syntax nat} ')')?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   674
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   675
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   676
  \begin{description}
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   677
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   678
  \item @{command (HOL) "recdef"} defines general well-founded
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   679
  recursive functions (using the TFL package), see also
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   680
  \cite{isabelle-HOL}.  The ``@{text "(permissive)"}'' option tells
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   681
  TFL to recover from failed proof attempts, returning unfinished
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   682
  results.  The @{text recdef_simp}, @{text recdef_cong}, and @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   683
  recdef_wf} hints refer to auxiliary rules to be used in the internal
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   684
  automated proof process of TFL.  Additional @{syntax clasimpmod}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   685
  declarations (cf.\ \secref{sec:clasimp}) may be given to tune the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   686
  context of the Simplifier (cf.\ \secref{sec:simplifier}) and
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   687
  Classical reasoner (cf.\ \secref{sec:classical}).
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   688
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   689
  \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   690
  proof for leftover termination condition number @{text i} (default
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   691
  1) as generated by a @{command (HOL) "recdef"} definition of
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   692
  constant @{text c}.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
   693
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   694
  Note that in most cases, @{command (HOL) "recdef"} is able to finish
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   695
  its internal proofs without manual intervention.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   696
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   697
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   698
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   699
  \medskip Hints for @{command (HOL) "recdef"} may be also declared
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   700
  globally, using the following attributes.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   701
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   702
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   703
    @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   704
    @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   705
    @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   706
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   707
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   708
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   709
    (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   710
      @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del')
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   711
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   712
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   713
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   714
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   715
section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   716
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   717
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   718
  An \textbf{inductive definition} specifies the least predicate (or
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   719
  set) @{text R} closed under given rules: applying a rule to elements
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   720
  of @{text R} yields a result within @{text R}.  For example, a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   721
  structural operational semantics is an inductive definition of an
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   722
  evaluation relation.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   723
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   724
  Dually, a \textbf{coinductive definition} specifies the greatest
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   725
  predicate~/ set @{text R} that is consistent with given rules: every
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   726
  element of @{text R} can be seen as arising by applying a rule to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   727
  elements of @{text R}.  An important example is using bisimulation
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   728
  relations to formalise equivalence of processes and infinite data
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   729
  structures.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   730
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   731
  \medskip The HOL package is related to the ZF one, which is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   732
  described in a separate paper,\footnote{It appeared in CADE
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   733
  \cite{paulson-CADE}; a longer version is distributed with Isabelle.}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   734
  which you should refer to in case of difficulties.  The package is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   735
  simpler than that of ZF thanks to implicit type-checking in HOL.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   736
  The types of the (co)inductive predicates (or sets) determine the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   737
  domain of the fixedpoint definition, and the package does not have
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   738
  to use inference rules for type-checking.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   739
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   740
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   741
    @{command_def (HOL) "inductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   742
    @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   743
    @{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   744
    @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   745
    @{attribute_def (HOL) mono} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   746
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   747
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   748
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   749
    (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   750
      @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set})
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   751
    @{syntax target}? @{syntax \"fixes\"} (@'for' @{syntax \"fixes\"})? \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   752
    (@'where' clauses)? (@'monos' @{syntax thmrefs})?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   753
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   754
    clauses: (@{syntax thmdecl}? @{syntax prop} + '|')
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   755
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   756
    @@{attribute (HOL) mono} (() | 'add' | 'del')
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   757
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   758
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   759
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   760
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   761
  \item @{command (HOL) "inductive"} and @{command (HOL)
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   762
  "coinductive"} define (co)inductive predicates from the
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   763
  introduction rules given in the @{keyword "where"} part.  The
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   764
  optional @{keyword "for"} part contains a list of parameters of the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   765
  (co)inductive predicates that remain fixed throughout the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   766
  definition.  The optional @{keyword "monos"} section contains
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   767
  \emph{monotonicity theorems}, which are required for each operator
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   768
  applied to a recursive set in the introduction rules.  There
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   769
  \emph{must} be a theorem of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"},
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   770
  for each premise @{text "M R\<^sub>i t"} in an introduction rule!
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   771
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   772
  \item @{command (HOL) "inductive_set"} and @{command (HOL)
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   773
  "coinductive_set"} are wrappers for to the previous commands,
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   774
  allowing the definition of (co)inductive sets.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   775
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   776
  \item @{attribute (HOL) mono} declares monotonicity rules.  These
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   777
  rule are involved in the automated monotonicity proof of @{command
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   778
  (HOL) "inductive"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   779
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   780
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   781
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   782
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   783
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   784
subsection {* Derived rules *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   785
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   786
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   787
  Each (co)inductive definition @{text R} adds definitions to the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   788
  theory and also proves some theorems:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   789
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   790
  \begin{description}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   791
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   792
  \item @{text R.intros} is the list of introduction rules as proven
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   793
  theorems, for the recursive predicates (or sets).  The rules are
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   794
  also available individually, using the names given them in the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   795
  theory file;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   796
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   797
  \item @{text R.cases} is the case analysis (or elimination) rule;
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   798
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   799
  \item @{text R.induct} or @{text R.coinduct} is the (co)induction
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   800
  rule.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   801
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   802
  \end{description}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   803
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   804
  When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   805
  defined simultaneously, the list of introduction rules is called
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   806
  @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   807
  called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   808
  of mutual induction rules is called @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   809
  "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   810
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   811
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   812
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   813
subsection {* Monotonicity theorems *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   814
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   815
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   816
  Each theory contains a default set of theorems that are used in
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   817
  monotonicity proofs.  New rules can be added to this set via the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   818
  @{attribute (HOL) mono} attribute.  The HOL theory @{text Inductive}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   819
  shows how this is done.  In general, the following monotonicity
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   820
  theorems may be added:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   821
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   822
  \begin{itemize}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   823
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   824
  \item Theorems of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"}, for proving
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   825
  monotonicity of inductive definitions whose introduction rules have
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   826
  premises involving terms such as @{text "M R\<^sub>i t"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   827
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   828
  \item Monotonicity theorems for logical operators, which are of the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   829
  general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}.  For example, in
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   830
  the case of the operator @{text "\<or>"}, the corresponding theorem is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   831
  \[
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   832
  \infer{@{text "P\<^sub>1 \<or> P\<^sub>2 \<longrightarrow> Q\<^sub>1 \<or> Q\<^sub>2"}}{@{text "P\<^sub>1 \<longrightarrow> Q\<^sub>1"} & @{text "P\<^sub>2 \<longrightarrow> Q\<^sub>2"}}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   833
  \]
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   834
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   835
  \item De Morgan style equations for reasoning about the ``polarity''
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   836
  of expressions, e.g.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   837
  \[
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   838
  @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   839
  @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   840
  \]
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   841
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   842
  \item Equations for reducing complex operators to more primitive
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   843
  ones whose monotonicity can easily be proved, e.g.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   844
  \[
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   845
  @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   846
  @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   847
  \]
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   848
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   849
  \end{itemize}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   850
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   851
  %FIXME: Example of an inductive definition
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   852
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   853
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   854
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   855
section {* Arithmetic proof support *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   856
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   857
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   858
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   859
    @{method_def (HOL) arith} & : & @{text method} \\
30863
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   860
    @{attribute_def (HOL) arith} & : & @{text attribute} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   861
    @{attribute_def (HOL) arith_split} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   862
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   863
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   864
  The @{method (HOL) arith} method decides linear arithmetic problems
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   865
  (on types @{text nat}, @{text int}, @{text real}).  Any current
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   866
  facts are inserted into the goal before running the procedure.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   867
30863
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   868
  The @{attribute (HOL) arith} attribute declares facts that are
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   869
  always supplied to the arithmetic provers implicitly.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   870
30863
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   871
  The @{attribute (HOL) arith_split} attribute declares case split
30865
5106e13d400f fixed formal markup;
wenzelm
parents: 30863
diff changeset
   872
  rules to be expanded before @{method (HOL) arith} is invoked.
30863
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   873
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   874
  Note that a simpler (but faster) arithmetic prover is
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   875
  already invoked by the Simplifier.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   876
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   877
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   878
30169
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   879
section {* Intuitionistic proof search *}
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   880
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   881
text {*
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   882
  \begin{matharray}{rcl}
30171
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   883
    @{method_def (HOL) iprover} & : & @{text method} \\
30169
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   884
  \end{matharray}
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   885
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   886
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   887
    @@{method (HOL) iprover} ( @{syntax rulemod} * )
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   888
  "}
30169
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   889
30171
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   890
  The @{method (HOL) iprover} method performs intuitionistic proof
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   891
  search, depending on specifically declared rules from the context,
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   892
  or given as explicit arguments.  Chained facts are inserted into the
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 35351
diff changeset
   893
  goal before commencing proof search.
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 35351
diff changeset
   894
30169
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   895
  Rules need to be classified as @{attribute (Pure) intro},
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   896
  @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   897
  ``@{text "!"}'' indicator refers to ``safe'' rules, which may be
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   898
  applied aggressively (without considering back-tracking later).
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   899
  Rules declared with ``@{text "?"}'' are ignored in proof search (the
42626
6ac8c55c657e eliminated some duplicate "def" positions;
wenzelm
parents: 42617
diff changeset
   900
  single-step @{method (Pure) rule} method still observes these).  An
30169
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   901
  explicit weight annotation may be given as well; otherwise the
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   902
  number of rule premises will be taken into account here.
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   903
*}
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   904
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   905
30171
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   906
section {* Coherent Logic *}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   907
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   908
text {*
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   909
  \begin{matharray}{rcl}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   910
    @{method_def (HOL) "coherent"} & : & @{text method} \\
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   911
  \end{matharray}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   912
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   913
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   914
    @@{method (HOL) coherent} @{syntax thmrefs}?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   915
  "}
30171
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   916
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   917
  The @{method (HOL) coherent} method solves problems of
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   918
  \emph{Coherent Logic} \cite{Bezem-Coquand:2005}, which covers
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   919
  applications in confluence theory, lattice theory and projective
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40711
diff changeset
   920
  geometry.  See @{file "~~/src/HOL/ex/Coherent.thy"} for some
30171
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   921
  examples.
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   922
*}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   923
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   924
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   925
section {* Proving propositions *}
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   926
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   927
text {*
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   928
  In addition to the standard proof methods, a number of diagnosis
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   929
  tools search for proofs and provide an Isar proof snippet on success.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   930
  These tools are available via the following commands.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   931
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   932
  \begin{matharray}{rcl}
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   933
    @{command_def (HOL) "solve_direct"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
43040
665623e695ea document new "try"
blanchet
parents: 43019
diff changeset
   934
    @{command_def (HOL) "try"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
43016
42330f25142c renamed "try" "try_methods"
blanchet
parents: 42705
diff changeset
   935
    @{command_def (HOL) "try_methods"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   936
    @{command_def (HOL) "sledgehammer"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   937
    @{command_def (HOL) "sledgehammer_params"} & : & @{text "theory \<rightarrow> theory"}
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   938
  \end{matharray}
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   939
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   940
  @{rail "
43040
665623e695ea document new "try"
blanchet
parents: 43019
diff changeset
   941
    @@{command (HOL) try}
665623e695ea document new "try"
blanchet
parents: 43019
diff changeset
   942
    ;
665623e695ea document new "try"
blanchet
parents: 43019
diff changeset
   943
43016
42330f25142c renamed "try" "try_methods"
blanchet
parents: 42705
diff changeset
   944
    @@{command (HOL) try_methods} ( ( ( 'simp' | 'intro' | 'elim' | 'dest' ) ':' @{syntax thmrefs} ) + ) ?
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   945
      @{syntax nat}?
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   946
    ;
43040
665623e695ea document new "try"
blanchet
parents: 43019
diff changeset
   947
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   948
    @@{command (HOL) sledgehammer} ( '[' args ']' )? facts? @{syntax nat}?
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   949
    ;
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   950
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   951
    @@{command (HOL) sledgehammer_params} ( ( '[' args ']' ) ? )
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   952
    ;
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   953
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   954
    args: ( @{syntax name} '=' value + ',' )
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   955
    ;
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   956
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
   957
    facts: '(' ( ( ( ( 'add' | 'del' ) ':' ) ? @{syntax thmrefs} ) + ) ? ')'
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   958
    ;
43019
619f16bf2150 removed FIXME: "try_methods" doesn't take a clasimpmod, because it needs to generate Isar proof text in case of success and hence needs total control over its arguments
blanchet
parents: 43016
diff changeset
   959
  "} % FIXME check args "value"
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   960
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   961
  \begin{description}
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   962
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   963
  \item @{command (HOL) "solve_direct"} checks whether the current subgoals can
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   964
    be solved directly by an existing theorem. Duplicate lemmas can be detected
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   965
    in this way.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   966
43016
42330f25142c renamed "try" "try_methods"
blanchet
parents: 42705
diff changeset
   967
  \item @{command (HOL) "try_methods"} attempts to prove a subgoal using a combination
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   968
    of standard proof methods (@{text auto}, @{text simp}, @{text blast}, etc.).
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   969
    Additional facts supplied via @{text "simp:"}, @{text "intro:"},
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   970
    @{text "elim:"}, and @{text "dest:"} are passed to the appropriate proof
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   971
    methods.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   972
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   973
  \item @{command (HOL) "sledgehammer"} attempts to prove a subgoal using external
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   974
    automatic provers (resolution provers and SMT solvers). See the Sledgehammer
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   975
    manual \cite{isabelle-sledgehammer} for details.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   976
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   977
  \item @{command (HOL) "sledgehammer_params"} changes
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   978
    @{command (HOL) "sledgehammer"} configuration options persistently.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   979
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   980
  \end{description}
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   981
*}
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   982
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   983
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   984
section {* Checking and refuting propositions *}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   985
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   986
text {*
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   987
  Identifying incorrect propositions usually involves evaluation of
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   988
  particular assignments and systematic counterexample search.  This
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   989
  is supported by the following commands.
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   990
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   991
  \begin{matharray}{rcl}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   992
    @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   993
    @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   994
    @{command_def (HOL) "refute"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   995
    @{command_def (HOL) "nitpick"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   996
    @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \<rightarrow> theory"} \\
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   997
    @{command_def (HOL) "refute_params"} & : & @{text "theory \<rightarrow> theory"} \\
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
   998
    @{command_def (HOL) "nitpick_params"} & : & @{text "theory \<rightarrow> theory"}
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   999
  \end{matharray}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1000
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1001
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1002
    @@{command (HOL) value} ( '[' name ']' )? modes? @{syntax term}
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1003
    ;
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1004
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1005
    (@@{command (HOL) quickcheck} | @@{command (HOL) refute} | @@{command (HOL) nitpick})
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1006
      ( '[' args ']' )? @{syntax nat}?
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1007
    ;
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1008
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1009
    (@@{command (HOL) quickcheck_params} | @@{command (HOL) refute_params} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1010
      @@{command (HOL) nitpick_params}) ( '[' args ']' )?
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1011
    ;
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1012
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1013
    modes: '(' (@{syntax name} +) ')'
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1014
    ;
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1015
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1016
    args: ( @{syntax name} '=' value + ',' )
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1017
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1018
  "} % FIXME check "value"
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1019
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1020
  \begin{description}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1021
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1022
  \item @{command (HOL) "value"}~@{text t} evaluates and prints a
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1023
    term; optionally @{text modes} can be specified, which are
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1024
    appended to the current print mode (see also \cite{isabelle-ref}).
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1025
    Internally, the evaluation is performed by registered evaluators,
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1026
    which are invoked sequentially until a result is returned.
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1027
    Alternatively a specific evaluator can be selected using square
37444
2e7e7ff21e25 added simp evaluator
haftmann
parents: 37422
diff changeset
  1028
    brackets; typical evaluators use the current set of code equations
2e7e7ff21e25 added simp evaluator
haftmann
parents: 37422
diff changeset
  1029
    to normalize and include @{text simp} for fully symbolic evaluation
2e7e7ff21e25 added simp evaluator
haftmann
parents: 37422
diff changeset
  1030
    using the simplifier, @{text nbe} for \emph{normalization by evaluation}
2e7e7ff21e25 added simp evaluator
haftmann
parents: 37422
diff changeset
  1031
    and \emph{code} for code generation in SML.
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1032
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1033
  \item @{command (HOL) "quickcheck"} tests the current goal for
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1034
    counterexamples using a series of assignments for its
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1035
    free variables; by default the first subgoal is tested, an other
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1036
    can be selected explicitly using an optional goal index.
40918
7351c6afb348 explaining quickcheck testers in the documentation
bulwahn
parents: 40800
diff changeset
  1037
    Assignments can be chosen exhausting the search space upto a given
7351c6afb348 explaining quickcheck testers in the documentation
bulwahn
parents: 40800
diff changeset
  1038
    size or using a fixed number of random assignments in the search space.
7351c6afb348 explaining quickcheck testers in the documentation
bulwahn
parents: 40800
diff changeset
  1039
    By default, quickcheck uses exhaustive testing.
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1040
    A number of configuration options are supported for
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1041
    @{command (HOL) "quickcheck"}, notably:
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1042
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1043
    \begin{description}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1044
40918
7351c6afb348 explaining quickcheck testers in the documentation
bulwahn
parents: 40800
diff changeset
  1045
    \item[@{text tester}] specifies how to explore the search space
7351c6afb348 explaining quickcheck testers in the documentation
bulwahn
parents: 40800
diff changeset
  1046
      (e.g. exhaustive or random).
7351c6afb348 explaining quickcheck testers in the documentation
bulwahn
parents: 40800
diff changeset
  1047
      An unknown configuration option is treated as an argument to tester,
7351c6afb348 explaining quickcheck testers in the documentation
bulwahn
parents: 40800
diff changeset
  1048
      making @{text "tester ="} optional.
40254
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1049
    \item[@{text size}] specifies the maximum size of the search space
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1050
    for assignment values.
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1051
42092
f07b373f25d3 adding documentation about the eval option in quickcheck
bulwahn
parents: 41846
diff changeset
  1052
    \item[@{text eval}] takes a term or a list of terms and evaluates
f07b373f25d3 adding documentation about the eval option in quickcheck
bulwahn
parents: 41846
diff changeset
  1053
      these terms under the variable assignment found by quickcheck.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
  1054
40254
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1055
    \item[@{text iterations}] sets how many sets of assignments are
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1056
    generated for each particular size.
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1057
40254
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1058
    \item[@{text no_assms}] specifies whether assumptions in
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1059
    structured proofs should be ignored.
35331
450ab945c451 document Quickcheck's "no_assms" option
blanchet
parents: 34172
diff changeset
  1060
40254
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1061
    \item[@{text timeout}] sets the time limit in seconds.
40245
59f011c1877a updating documentation on quickcheck in the Isar reference
bulwahn
parents: 40171
diff changeset
  1062
40254
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1063
    \item[@{text default_type}] sets the type(s) generally used to
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1064
    instantiate type variables.
40245
59f011c1877a updating documentation on quickcheck in the Isar reference
bulwahn
parents: 40171
diff changeset
  1065
40254
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1066
    \item[@{text report}] if set quickcheck reports how many tests
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1067
    fulfilled the preconditions.
40245
59f011c1877a updating documentation on quickcheck in the Isar reference
bulwahn
parents: 40171
diff changeset
  1068
40254
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1069
    \item[@{text quiet}] if not set quickcheck informs about the
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1070
    current size for assignment values.
40245
59f011c1877a updating documentation on quickcheck in the Isar reference
bulwahn
parents: 40171
diff changeset
  1071
40254
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1072
    \item[@{text expect}] can be used to check if the user's
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1073
    expectation was met (@{text no_expectation}, @{text
6d1ebaa7a4ba proper markup of formal text;
wenzelm
parents: 40245
diff changeset
  1074
    no_counterexample}, or @{text counterexample}).
40245
59f011c1877a updating documentation on quickcheck in the Isar reference
bulwahn
parents: 40171
diff changeset
  1075
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1076
    \end{description}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1077
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1078
    These option can be given within square brackets.
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1079
42215
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1080
  \item @{command (HOL) "quickcheck_params"} changes
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1081
    @{command (HOL) "quickcheck"} configuration options persistently.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1082
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1083
  \item @{command (HOL) "refute"} tests the current goal for
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1084
    counterexamples using a reduction to SAT. The following configuration
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1085
    options are supported:
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1086
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1087
    \begin{description}
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1088
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1089
    \item[@{text minsize}] specifies the minimum size (cardinality) of the
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1090
      models to search for.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1091
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1092
    \item[@{text maxsize}] specifies the maximum size (cardinality) of the
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1093
      models to search for. Nonpositive values mean $\infty$.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1094
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1095
    \item[@{text maxvars}] specifies the maximum number of Boolean variables
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1096
    to use when transforming the term into a propositional formula.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1097
    Nonpositive values mean $\infty$.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1098
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1099
    \item[@{text satsolver}] specifies the SAT solver to use.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1100
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1101
    \item[@{text no_assms}] specifies whether assumptions in
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1102
    structured proofs should be ignored.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1103
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1104
    \item[@{text maxtime}] sets the time limit in seconds.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1105
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1106
    \item[@{text expect}] can be used to check if the user's
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1107
    expectation was met (@{text genuine}, @{text potential},
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1108
    @{text none}, or @{text unknown}).
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1109
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1110
    \end{description}
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1111
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1112
    These option can be given within square brackets.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1113
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1114
  \item @{command (HOL) "refute_params"} changes
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1115
    @{command (HOL) "refute"} configuration options persistently.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1116
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1117
  \item @{command (HOL) "nitpick"} tests the current goal for counterexamples
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1118
    using a reduction to first-order relational logic. See the Nitpick manual
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1119
    \cite{isabelle-nitpick} for details.
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1120
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1121
  \item @{command (HOL) "nitpick_params"} changes
de9d43c427ae document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents: 42123
diff changeset
  1122
    @{command (HOL) "nitpick"} configuration options persistently.
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1123
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1124
  \end{description}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1125
*}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1126
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
  1127
28752
wenzelm
parents: 28687
diff changeset
  1128
section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1129
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1130
text {*
27123
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1131
  The following tools of Isabelle/HOL support cases analysis and
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1132
  induction in unstructured tactic scripts; see also
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1133
  \secref{sec:cases-induct} for proper Isar versions of similar ideas.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1134
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1135
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1136
    @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1137
    @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1138
    @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1139
    @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1140
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1141
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1142
  @{rail "
42705
528a2ba8fa74 tuned some syntax names;
wenzelm
parents: 42704
diff changeset
  1143
    @@{method (HOL) case_tac} @{syntax goal_spec}? @{syntax term} rule?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1144
    ;
42705
528a2ba8fa74 tuned some syntax names;
wenzelm
parents: 42704
diff changeset
  1145
    @@{method (HOL) induct_tac} @{syntax goal_spec}? (@{syntax insts} * @'and') rule?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1146
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1147
    @@{method (HOL) ind_cases} (@{syntax prop}+) (@'for' (@{syntax name}+))?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1148
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1149
    @@{command (HOL) inductive_cases} (@{syntax thmdecl}? (@{syntax prop}+) + @'and')
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1150
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1151
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1152
    rule: 'rule' ':' @{syntax thmref}
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1153
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1154
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1155
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1156
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1157
  \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1158
  to reason about inductive types.  Rules are selected according to
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1159
  the declarations by the @{attribute cases} and @{attribute induct}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1160
  attributes, cf.\ \secref{sec:cases-induct}.  The @{command (HOL)
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1161
  datatype} package already takes care of this.
27123
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1162
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1163
  These unstructured tactics feature both goal addressing and dynamic
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1164
  instantiation.  Note that named rule cases are \emph{not} provided
27123
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1165
  as would be by the proper @{method cases} and @{method induct} proof
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1166
  methods (see \secref{sec:cases-induct}).  Unlike the @{method
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1167
  induct} method, @{method induct_tac} does not handle structured rule
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1168
  statements, only the compact object-logic conclusion of the subgoal
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1169
  being addressed.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
  1170
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1171
  \item @{method (HOL) ind_cases} and @{command (HOL)
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1172
  "inductive_cases"} provide an interface to the internal @{ML_text
26860
7c749112261c replaced some latex macros by antiquotations;
wenzelm
parents: 26852
diff changeset
  1173
  mk_cases} operation.  Rules are simplified in an unrestricted
7c749112261c replaced some latex macros by antiquotations;
wenzelm
parents: 26852
diff changeset
  1174
  forward manner.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1175
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1176
  While @{method (HOL) ind_cases} is a proof method to apply the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1177
  result immediately as elimination rules, @{command (HOL)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1178
  "inductive_cases"} provides case split theorems at the theory level
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1179
  for later use.  The @{keyword "for"} argument of the @{method (HOL)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1180
  ind_cases} method allows to specify a list of variables that should
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1181
  be generalized before applying the resulting rule.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1182
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1183
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1184
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1185
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1186
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1187
section {* Executable code *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1188
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1189
text {* For validation purposes, it is often useful to \emph{execute}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1190
  specifications.  In principle, execution could be simulated by
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1191
  Isabelle's inference kernel, i.e. by a combination of resolution and
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1192
  simplification.  Unfortunately, this approach is rather inefficient.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1193
  A more efficient way of executing specifications is to translate
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1194
  them into a functional programming language such as ML.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1195
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1196
  Isabelle provides two generic frameworks to support code generation
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1197
  from executable specifications.  Isabelle/HOL instantiates these
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1198
  mechanisms in a way that is amenable to end-user applications.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1199
*}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1200
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1201
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1202
subsection {* The new code generator (F. Haftmann) *}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1203
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1204
text {* This framework generates code from functional programs
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1205
  (including overloading using type classes) to SML \cite{SML}, OCaml
38814
4d575fbfc920 official support for Scala
haftmann
parents: 38462
diff changeset
  1206
  \cite{OCaml}, Haskell \cite{haskell-revised-report} and Scala
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1207
  \cite{scala-overview-tech-report}.  Conceptually, code generation is
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1208
  split up in three steps: \emph{selection} of code theorems,
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1209
  \emph{translation} into an abstract executable view and
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1210
  \emph{serialization} to a specific \emph{target language}.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1211
  Inductive specifications can be executed using the predicate
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1212
  compiler which operates within HOL.  See \cite{isabelle-codegen} for
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1213
  an introduction.
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1214
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1215
  \begin{matharray}{rcl}
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1216
    @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1217
    @{attribute_def (HOL) code} & : & @{text attribute} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1218
    @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1219
    @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1220
    @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1221
    @{attribute_def (HOL) code_inline} & : & @{text attribute} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1222
    @{attribute_def (HOL) code_post} & : & @{text attribute} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1223
    @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1224
    @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1225
    @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1226
    @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1227
    @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1228
    @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1229
    @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1230
    @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1231
    @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1232
    @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1233
    @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1234
    @{command_def (HOL) "code_reflect"} & : & @{text "theory \<rightarrow> theory"}
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1235
  \end{matharray}
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1236
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1237
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1238
    @@{command (HOL) export_code} ( constexpr + ) \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1239
       ( ( @'in' target ( @'module_name' @{syntax string} ) ? \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1240
        ( @'file' ( @{syntax string} | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1241
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1242
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1243
    const: @{syntax term}
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1244
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1245
40711
81bc73585eec globbing constant expressions use more idiomatic underscore rather than star
haftmann
parents: 40709
diff changeset
  1246
    constexpr: ( const | 'name._' | '_' )
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1247
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1248
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1249
    typeconstructor: @{syntax nameref}
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1250
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1251
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1252
    class: @{syntax nameref}
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1253
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1254
38814
4d575fbfc920 official support for Scala
haftmann
parents: 38462
diff changeset
  1255
    target: 'SML' | 'OCaml' | 'Haskell' | 'Scala'
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1256
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1257
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1258
    @@{attribute (HOL) code} ( 'del' | 'abstype' | 'abstract' )?
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1259
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1260
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1261
    @@{command (HOL) code_abort} ( const + )
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1262
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1263
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1264
    @@{command (HOL) code_datatype} ( const + )
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1265
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1266
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1267
    @@{attribute (HOL) code_inline} ( 'del' ) ?
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1268
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1269
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1270
    @@{attribute (HOL) code_post} ( 'del' ) ?
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1271
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1272
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1273
    @@{command (HOL) code_thms} ( constexpr + ) ?
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1274
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1275
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1276
    @@{command (HOL) code_deps} ( constexpr + ) ?
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1277
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1278
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1279
    @@{command (HOL) code_const} (const + @'and') \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1280
      ( ( '(' target ( syntax ? + @'and' ) ')' ) + )
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1281
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1282
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1283
    @@{command (HOL) code_type} (typeconstructor + @'and') \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1284
      ( ( '(' target ( syntax ? + @'and' ) ')' ) + )
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1285
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1286
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1287
    @@{command (HOL) code_class} (class + @'and') \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1288
      ( ( '(' target \\ ( @{syntax string} ? + @'and' ) ')' ) + )
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1289
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1290
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1291
    @@{command (HOL) code_instance} (( typeconstructor '::' class ) + @'and') \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1292
      ( ( '(' target ( '-' ? + @'and' ) ')' ) + )
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1293
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1294
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1295
    @@{command (HOL) code_reserved} target ( @{syntax string} + )
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1296
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1297
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1298
    @@{command (HOL) code_monad} const const target
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1299
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1300
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1301
    @@{command (HOL) code_include} target ( @{syntax string} ( @{syntax string} | '-') )
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1302
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1303
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1304
    @@{command (HOL) code_modulename} target ( ( @{syntax string} @{syntax string} ) + )
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1305
    ;
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1306
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1307
    @@{command (HOL) code_reflect} @{syntax string} \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1308
      ( @'datatypes' ( @{syntax string} '=' ( '_' | ( @{syntax string} + '|' ) + @'and' ) ) ) ? \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1309
      ( @'functions' ( @{syntax string} + ) ) ? ( @'file' @{syntax string} ) ?
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1310
    ;
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1311
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1312
    syntax: @{syntax string} | ( @'infix' | @'infixl' | @'infixr' ) @{syntax nat} @{syntax string}
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1313
  "}
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1314
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1315
  \begin{description}
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1316
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1317
  \item @{command (HOL) "export_code"} generates code for a given list
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1318
  of constants in the specified target language(s).  If no
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1319
  serialization instruction is given, only abstract code is generated
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1320
  internally.
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1321
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1322
  Constants may be specified by giving them literally, referring to
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1323
  all executable contants within a certain theory by giving @{text
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1324
  "name.*"}, or referring to \emph{all} executable constants currently
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1325
  available by giving @{text "*"}.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1326
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1327
  By default, for each involved theory one corresponding name space
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1328
  module is generated.  Alternativly, a module name may be specified
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1329
  after the @{keyword "module_name"} keyword; then \emph{all} code is
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1330
  placed in this module.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1331
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1332
  For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1333
  refers to a single file; for \emph{Haskell}, it refers to a whole
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1334
  directory, where code is generated in multiple files reflecting the
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1335
  module hierarchy.  Omitting the file specification denotes standard
37749
c7e15d59c58d updated documentation
haftmann
parents: 37444
diff changeset
  1336
  output.
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1337
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1338
  Serializers take an optional list of arguments in parentheses.  For
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1339
  \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1340
  explicit module signatures.
42123
c407078c0d47 updated generated file;
wenzelm
parents: 42092
diff changeset
  1341
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1342
  For \emph{Haskell} a module name prefix may be given using the
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1343
  ``@{text "root:"}'' argument; ``@{text string_classes}'' adds a
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1344
  ``@{verbatim "deriving (Read, Show)"}'' clause to each appropriate
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1345
  datatype declaration.
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1346
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1347
  \item @{attribute (HOL) code} explicitly selects (or with option
38462
34d3de1254cd formally document `code abstype` and `code abstract` attributes
haftmann
parents: 37820
diff changeset
  1348
  ``@{text "del"}'' deselects) a code equation for code generation.
34d3de1254cd formally document `code abstype` and `code abstract` attributes
haftmann
parents: 37820
diff changeset
  1349
  Usually packages introducing code equations provide a reasonable
34d3de1254cd formally document `code abstype` and `code abstract` attributes
haftmann
parents: 37820
diff changeset
  1350
  default setup for selection.  Variants @{text "code abstype"} and
34d3de1254cd formally document `code abstype` and `code abstract` attributes
haftmann
parents: 37820
diff changeset
  1351
  @{text "code abstract"} declare abstract datatype certificates or
34d3de1254cd formally document `code abstype` and `code abstract` attributes
haftmann
parents: 37820
diff changeset
  1352
  code equations on abstract datatype representations respectively.
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1353
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1354
  \item @{command (HOL) "code_abort"} declares constants which are not
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1355
  required to have a definition by means of code equations; if needed
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1356
  these are implemented by program abort instead.
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1357
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1358
  \item @{command (HOL) "code_datatype"} specifies a constructor set
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1359
  for a logical type.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1360
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1361
  \item @{command (HOL) "print_codesetup"} gives an overview on
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1362
  selected code equations and code generator datatypes.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1363
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1364
  \item @{attribute (HOL) code_inline} declares (or with option
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1365
  ``@{text "del"}'' removes) inlining theorems which are applied as
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1366
  rewrite rules to any code equation during preprocessing.
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1367
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1368
  \item @{attribute (HOL) code_post} declares (or with option ``@{text
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1369
  "del"}'' removes) theorems which are applied as rewrite rules to any
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1370
  result of an evaluation.
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1371
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1372
  \item @{command (HOL) "print_codeproc"} prints the setup of the code
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1373
  generator preprocessor.
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1374
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1375
  \item @{command (HOL) "code_thms"} prints a list of theorems
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1376
  representing the corresponding program containing all given
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1377
  constants after preprocessing.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1378
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1379
  \item @{command (HOL) "code_deps"} visualizes dependencies of
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1380
  theorems representing the corresponding program containing all given
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1381
  constants after preprocessing.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1382
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1383
  \item @{command (HOL) "code_const"} associates a list of constants
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1384
  with target-specific serializations; omitting a serialization
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1385
  deletes an existing serialization.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1386
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1387
  \item @{command (HOL) "code_type"} associates a list of type
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1388
  constructors with target-specific serializations; omitting a
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1389
  serialization deletes an existing serialization.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1390
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1391
  \item @{command (HOL) "code_class"} associates a list of classes
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1392
  with target-specific class names; omitting a serialization deletes
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1393
  an existing serialization.  This applies only to \emph{Haskell}.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1394
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1395
  \item @{command (HOL) "code_instance"} declares a list of type
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1396
  constructor / class instance relations as ``already present'' for a
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1397
  given target.  Omitting a ``@{text "-"}'' deletes an existing
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1398
  ``already present'' declaration.  This applies only to
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1399
  \emph{Haskell}.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1400
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1401
  \item @{command (HOL) "code_reserved"} declares a list of names as
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1402
  reserved for a given target, preventing it to be shadowed by any
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1403
  generated code.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1404
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1405
  \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1406
  to generate monadic code for Haskell.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1407
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1408
  \item @{command (HOL) "code_include"} adds arbitrary named content
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1409
  (``include'') to generated code.  A ``@{text "-"}'' as last argument
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1410
  will remove an already added ``include''.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1411
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1412
  \item @{command (HOL) "code_modulename"} declares aliasings from one
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1413
  module name onto another.
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1414
39608
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1415
  \item @{command (HOL) "code_reflect"} without a ``@{text "file"}''
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1416
  argument compiles code into the system runtime environment and
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1417
  modifies the code generator setup that future invocations of system
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1418
  runtime code generation referring to one of the ``@{text
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1419
  "datatypes"}'' or ``@{text "functions"}'' entities use these precompiled
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1420
  entities.  With a ``@{text "file"}'' argument, the corresponding code
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1421
  is generated into that specified file without modifying the code
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1422
  generator setup.
76bc7e4999f8 formal syntax diagram for code_reflect
haftmann
parents: 38814
diff changeset
  1423
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1424
  \end{description}
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1425
*}
37422
6d19e4e6ebf5 tuned internal order
haftmann
parents: 37379
diff changeset
  1426
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1427
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1428
subsection {* The old code generator (S. Berghofer) *}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1429
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1430
text {* This framework generates code from both functional and
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1431
  relational programs to SML, as explained below.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1432
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1433
  \begin{matharray}{rcl}
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1434
    @{command_def "code_module"} & : & @{text "theory \<rightarrow> theory"} \\
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1435
    @{command_def "code_library"} & : & @{text "theory \<rightarrow> theory"} \\
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1436
    @{command_def "consts_code"} & : & @{text "theory \<rightarrow> theory"} \\
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1437
    @{command_def "types_code"} & : & @{text "theory \<rightarrow> theory"} \\
42626
6ac8c55c657e eliminated some duplicate "def" positions;
wenzelm
parents: 42617
diff changeset
  1438
    @{attribute_def code} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1439
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1440
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1441
  @{rail "
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1442
  ( @@{command code_module} | @@{command code_library} ) modespec? @{syntax name}? \\
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1443
    ( @'file' name ) ? ( @'imports' ( @{syntax name} + ) ) ? \\
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1444
    @'contains' ( ( @{syntax name} '=' @{syntax term} ) + | @{syntax term} + )
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1445
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1446
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1447
  modespec: '(' ( @{syntax name} * ) ')'
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1448
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1449
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1450
  @@{command (HOL) consts_code} (codespec +)
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1451
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1452
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1453
  codespec: const template attachment ?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1454
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1455
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1456
  @@{command (HOL) types_code} (tycodespec +)
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1457
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1458
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1459
  tycodespec: @{syntax name} template attachment ?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1460
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1461
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1462
  const: @{syntax term}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1463
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1464
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1465
  template: '(' @{syntax string} ')'
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1466
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1467
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1468
  attachment: 'attach' modespec? '{' @{syntax text} '}'
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1469
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1470
42626
6ac8c55c657e eliminated some duplicate "def" positions;
wenzelm
parents: 42617
diff changeset
  1471
  @@{attribute code} name?
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1472
  "}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1473
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1474
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1475
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1476
subsubsection {* Invoking the code generator *}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1477
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1478
text {* The code generator is invoked via the @{command code_module}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1479
  and @{command code_library} commands, which correspond to
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1480
  \emph{incremental} and \emph{modular} code generation, respectively.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1481
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1482
  \begin{description}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1483
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1484
  \item [Modular] For each theory, an ML structure is generated,
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1485
  containing the code generated from the constants defined in this
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1486
  theory.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1487
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1488
  \item [Incremental] All the generated code is emitted into the same
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1489
  structure.  This structure may import code from previously generated
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1490
  structures, which can be specified via @{keyword "imports"}.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1491
  Moreover, the generated structure may also be referred to in later
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1492
  invocations of the code generator.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1493
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1494
  \end{description}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1495
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1496
  After the @{command code_module} and @{command code_library}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1497
  keywords, the user may specify an optional list of ``modes'' in
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1498
  parentheses. These can be used to instruct the code generator to
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1499
  emit additional code for special purposes, e.g.\ functions for
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1500
  converting elements of generated datatypes to Isabelle terms, or
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1501
  test data generators. The list of modes is followed by a module
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1502
  name.  The module name is optional for modular code generation, but
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1503
  must be specified for incremental code generation.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1504
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1505
  The code can either be written to a file, in which case a file name
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1506
  has to be specified after the @{keyword "file"} keyword, or be loaded
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1507
  directly into Isabelle's ML environment. In the latter case, the
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1508
  @{command ML} theory command can be used to inspect the results
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1509
  interactively, for example.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1510
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1511
  The terms from which to generate code can be specified after the
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1512
  @{keyword "contains"} keyword, either as a list of bindings, or just
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1513
  as a list of terms. In the latter case, the code generator just
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1514
  produces code for all constants and types occuring in the term, but
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1515
  does not bind the compiled terms to ML identifiers.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1516
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1517
  Here is an example:
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1518
*}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1519
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1520
code_module Test
42652
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1521
contains test = "foldl op + (0 :: int) [1, 2, 3, 4, 5]"
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1522
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1523
text {* \noindent This binds the result of compiling the given term to
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1524
  the ML identifier @{ML Test.test}.  *}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1525
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1526
ML {* @{assert} (Test.test = 15) *}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1527
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1528
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1529
subsubsection {* Configuring the code generator *}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1530
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1531
text {* When generating code for a complex term, the code generator
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1532
  recursively calls itself for all subterms.  When it arrives at a
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1533
  constant, the default strategy of the code generator is to look up
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1534
  its definition and try to generate code for it.  Constants which
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1535
  have no definitions that are immediately executable, may be
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1536
  associated with a piece of ML code manually using the @{command_ref
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1537
  consts_code} command.  It takes a list whose elements consist of a
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1538
  constant (given in usual term syntax -- an explicit type constraint
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1539
  accounts for overloading), and a mixfix template describing the ML
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1540
  code. The latter is very much the same as the mixfix templates used
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1541
  when declaring new constants.  The most notable difference is that
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1542
  terms may be included in the ML template using antiquotation
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1543
  brackets @{verbatim "{"}@{verbatim "*"}~@{text "..."}~@{verbatim
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1544
  "*"}@{verbatim "}"}.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1545
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1546
  A similar mechanism is available for types: @{command_ref
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1547
  types_code} associates type constructors with specific ML code.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1548
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1549
  For example, the following declarations copied from @{file
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1550
  "~~/src/HOL/Product_Type.thy"} describe how the product type of
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1551
  Isabelle/HOL should be compiled to ML.  *}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1552
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1553
typedecl ('a, 'b) prod
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1554
consts Pair :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) prod"
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1555
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1556
types_code prod  ("(_ */ _)")
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1557
consts_code Pair  ("(_,/ _)")
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1558
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1559
text {* Sometimes, the code associated with a constant or type may
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1560
  need to refer to auxiliary functions, which have to be emitted when
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1561
  the constant is used. Code for such auxiliary functions can be
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1562
  declared using @{keyword "attach"}. For example, the @{const wfrec}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1563
  function can be implemented as follows:
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1564
*}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1565
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1566
consts_code wfrec  ("\<module>wfrec?")  (* FIXME !? *)
42652
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1567
attach {* fun wfrec f x = f (wfrec f) x *}
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1568
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1569
text {* If the code containing a call to @{const wfrec} resides in an
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1570
  ML structure different from the one containing the function
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1571
  definition attached to @{const wfrec}, the name of the ML structure
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1572
  (followed by a ``@{text "."}'')  is inserted in place of ``@{text
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1573
  "\<module>"}'' in the above template.  The ``@{text "?"}''  means that
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1574
  the code generator should ignore the first argument of @{const
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1575
  wfrec}, i.e.\ the termination relation, which is usually not
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1576
  executable.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1577
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1578
  \medskip Another possibility of configuring the code generator is to
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1579
  register theorems to be used for code generation. Theorems can be
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1580
  registered via the @{attribute code} attribute. It takes an optional
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1581
  name as an argument, which indicates the format of the
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1582
  theorem. Currently supported formats are equations (this is the
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1583
  default when no name is specified) and horn clauses (this is
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1584
  indicated by the name \texttt{ind}). The left-hand sides of
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1585
  equations may only contain constructors and distinct variables,
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1586
  whereas horn clauses must have the same format as introduction rules
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1587
  of inductive definitions.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1588
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1589
  The following example specifies three equations from which to
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1590
  generate code for @{term "op <"} on natural numbers (see also
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1591
  @{"file" "~~/src/HOL/Nat.thy"}).  *}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1592
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1593
lemma [code]: "(Suc m < Suc n) = (m < n)"
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1594
  and [code]: "((n::nat) < 0) = False"
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1595
  and [code]: "(0 < Suc n) = True" by simp_all
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1596
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1597
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1598
subsubsection {* Specific HOL code generators *}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1599
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1600
text {* The basic code generator framework offered by Isabelle/Pure
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1601
  has already been extended with additional code generators for
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1602
  specific HOL constructs. These include datatypes, recursive
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1603
  functions and inductive relations. The code generator for inductive
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1604
  relations can handle expressions of the form @{text "(t\<^sub>1, \<dots>, t\<^sub>n) \<in>
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1605
  r"}, where @{text "r"} is an inductively defined relation. If at
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1606
  least one of the @{text "t\<^sub>i"} is a dummy pattern ``@{text "_"}'',
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1607
  the above expression evaluates to a sequence of possible answers. If
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1608
  all of the @{text "t\<^sub>i"} are proper terms, the expression evaluates
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1609
  to a boolean value.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1610
42652
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1611
  The following example demonstrates this for beta-reduction on lambda
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1612
  terms (see also @{"file" "~~/src/HOL/Proofs/Lambda/Lambda.thy"}).
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1613
*}
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1614
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1615
datatype dB =
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1616
    Var nat
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1617
  | App dB dB  (infixl "\<degree>" 200)
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1618
  | Abs dB
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1619
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1620
primrec lift :: "dB \<Rightarrow> nat \<Rightarrow> dB"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1621
where
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1622
    "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1623
  | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1624
  | "lift (Abs s) k = Abs (lift s (k + 1))"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1625
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1626
primrec subst :: "dB \<Rightarrow> dB \<Rightarrow> nat \<Rightarrow> dB"  ("_[_'/_]" [300, 0, 0] 300)
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1627
where
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1628
    "(Var i)[s/k] =
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1629
      (if k < i then Var (i - 1) else if i = k then s else Var i)"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1630
  | "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1631
  | "(Abs t)[s/k] = Abs (t[lift s 0 / k+1])"
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1632
42652
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1633
inductive beta :: "dB \<Rightarrow> dB \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1634
where
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1635
    beta: "Abs s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1636
  | appL: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1637
  | appR: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1638
  | abs: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> Abs s \<rightarrow>\<^sub>\<beta> Abs t"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1639
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1640
code_module Test
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1641
contains
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1642
  test1 = "Abs (Var 0) \<degree> Var 0 \<rightarrow>\<^sub>\<beta> Var 0"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1643
  test2 = "Abs (Abs (Var 0 \<degree> Var 0) \<degree> (Abs (Var 0) \<degree> Var 0)) \<rightarrow>\<^sub>\<beta> _"
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1644
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1645
text {*
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1646
  In the above example, @{ML Test.test1} evaluates to a boolean,
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1647
  whereas @{ML Test.test2} is a lazy sequence whose elements can be
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1648
  inspected separately.
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1649
*}
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1650
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1651
ML {* @{assert} Test.test1 *}
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1652
ML {* val results = DSeq.list_of Test.test2 *}
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1653
ML {* @{assert} (length results = 2) *}
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1654
c963499143e5 reactivated codegen example based on Lambda.thy;
wenzelm
parents: 42651
diff changeset
  1655
text {*
42627
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1656
  \medskip The theory underlying the HOL code generator is described
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1657
  more detailed in \cite{Berghofer-Nipkow:2002}. More examples that
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1658
  illustrate the usage of the code generator can be found e.g.\ in
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1659
  @{"file" "~~/src/HOL/MicroJava/J/JListExample.thy"} and @{"file"
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1660
  "~~/src/HOL/MicroJava/JVM/JVMListExample.thy"}.
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1661
*}
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1662
8749742785b8 moved material about old codegen to isar-ref manual;
wenzelm
parents: 42626
diff changeset
  1663
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1664
section {* Definition by specification \label{sec:hol-specification} *}
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1665
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1666
text {*
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1667
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1668
    @{command_def (HOL) "specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1669
    @{command_def (HOL) "ax_specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1670
  \end{matharray}
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1671
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1672
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1673
  (@@{command (HOL) specification} | @@{command (HOL) ax_specification})
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1674
    '(' (decl +) ')' \\ (@{syntax thmdecl}? @{syntax prop} +)
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1675
  ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1676
  decl: ((@{syntax name} ':')? @{syntax term} '(' @'overloaded' ')'?)
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42215
diff changeset
  1677
  "}
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1678
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1679
  \begin{description}
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1680
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1681
  \item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1682
  goal stating the existence of terms with the properties specified to
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1683
  hold for the constants given in @{text decls}.  After finishing the
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1684
  proof, the theory will be augmented with definitions for the given
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1685
  constants, as well as with theorems stating the properties for these
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1686
  constants.
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1687
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1688
  \item @{command (HOL) "ax_specification"}~@{text "decls \<phi>"} sets up
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1689
  a goal stating the existence of terms with the properties specified
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1690
  to hold for the constants given in @{text decls}.  After finishing
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1691
  the proof, the theory will be augmented with axioms expressing the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1692
  properties given in the first place.
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1693
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1694
  \item @{text decl} declares a constant to be defined by the
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1695
  specification given.  The definition for the constant @{text c} is
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1696
  bound to the name @{text c_def} unless a theorem name is given in
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1697
  the declaration.  Overloaded constants should be declared as such.
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1698
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1699
  \end{description}
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1700
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1701
  Whether to use @{command (HOL) "specification"} or @{command (HOL)
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1702
  "ax_specification"} is to some extent a matter of style.  @{command
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1703
  (HOL) "specification"} introduces no new axioms, and so by
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1704
  construction cannot introduce inconsistencies, whereas @{command
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1705
  (HOL) "ax_specification"} does introduce axioms, but only after the
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1706
  user has explicitly proven it to be safe.  A practical issue must be
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1707
  considered, though: After introducing two constants with the same
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1708
  properties using @{command (HOL) "specification"}, one can prove
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1709
  that the two constants are, in fact, equal.  If this might be a
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1710
  problem, one should use @{command (HOL) "ax_specification"}.
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1711
*}
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1712
26840
ec46381f149d added logic-specific sessions;
wenzelm
parents:
diff changeset
  1713
end