doc-src/IsarRef/Thy/HOL_Specific.thy
author blanchet
Mon, 19 Apr 2010 18:14:45 +0200
changeset 36230 43d10a494c91
parent 36158 d2ad76e374d3
child 36453 2f383885d8f8
permissions -rw-r--r--
added warning about inconsistent context to Metis; it makes more sense here than in Sledgehammer, because Sledgehammer is unsound and there's no point in having people panicking about the consistency of their context when their context is in fact consistent
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26840
ec46381f149d added logic-specific sessions;
wenzelm
parents:
diff changeset
     1
theory HOL_Specific
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
     2
imports 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
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    14
  \begin{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    15
    'typedef' altname? abstype '=' repset
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    16
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    17
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    18
    altname: '(' (name | 'open' | 'open' name) ')'
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    19
    ;
35841
94f901e4969a allow sort constraints in HOL/typedef;
wenzelm
parents: 35744
diff changeset
    20
    abstype: typespecsorts mixfix?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    21
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    22
    repset: term ('morphisms' name name)?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    23
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    24
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    25
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    26
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    27
  
35744
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    28
  \item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"}
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    29
  axiomatizes a Gordon/HOL-style type definition in the background
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    30
  theory of the current context, depending on a non-emptiness result
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    31
  of the set @{text A} (which needs to be proven interactively).
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    32
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    33
  The raw type may not depend on parameters or assumptions of the
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    34
  context --- this is logically impossible in Isabelle/HOL --- but the
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    35
  non-emptiness property can be local, potentially resulting in
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    36
  multiple interpretations in target contexts.  Thus the established
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    37
  bijection between the representing set @{text A} and the new type
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    38
  @{text t} may semantically depend on local assumptions.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    39
  
35744
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    40
  By default, @{command (HOL) "typedef"} defines both a type @{text t}
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    41
  and a set (term constant) of the same name, unless an alternative
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    42
  base name is given in parentheses, or the ``@{text "(open)"}''
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    43
  declaration is used to suppress a separate constant definition
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    44
  altogether.  The injection from type to set is called @{text Rep_t},
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    45
  its inverse @{text Abs_t} --- this may be changed via an explicit
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    46
  @{keyword (HOL) "morphisms"} declaration.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    47
  
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    48
  Theorems @{text Rep_t}, @{text Rep_t_inverse}, and @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    49
  Abs_t_inverse} provide the most basic characterization as a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    50
  corresponding injection/surjection pair (in both directions).  Rules
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    51
  @{text Rep_t_inject} and @{text Abs_t_inject} provide a slightly
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    52
  more convenient view on the injectivity part, suitable for automated
26894
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26860
diff changeset
    53
  proof tools (e.g.\ in @{attribute simp} or @{attribute iff}
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26860
diff changeset
    54
  declarations).  Rules @{text Rep_t_cases}/@{text Rep_t_induct}, and
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26860
diff changeset
    55
  @{text Abs_t_cases}/@{text Abs_t_induct} provide alternative views
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26860
diff changeset
    56
  on surjectivity; these are already declared as set or type rules for
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    57
  the generic @{method cases} and @{method induct} methods.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    58
  
35744
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    59
  An alternative name for the set definition (and other derived
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    60
  entities) may be specified in parentheses; the default is to use
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
    61
  @{text t} as indicated before.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    62
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    63
  \end{description}
26849
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
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    67
section {* Adhoc tuples *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    68
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    69
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    70
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    71
    @{attribute (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    72
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    73
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    74
  \begin{rail}
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
    75
    'split\_format' ((( name * ) + 'and') | ('(' 'complete' ')'))
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    76
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    77
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    78
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    79
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    80
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    81
  \item @{attribute (HOL) split_format}~@{text "p\<^sub>1 \<dots> p\<^sub>m \<AND> \<dots>
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    82
  \<AND> q\<^sub>1 \<dots> q\<^sub>n"} puts expressions of low-level tuple types into
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    83
  canonical form as specified by the arguments given; the @{text i}-th
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    84
  collection of arguments refers to occurrences in premise @{text i}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    85
  of the rule.  The ``@{text "(complete)"}'' option causes \emph{all}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    86
  arguments in function applications to be represented canonically
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    87
  according to their tuple type structure.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    88
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    89
  Note that these operations tend to invent funny names for new local
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    90
  parameters to be introduced.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    91
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
    92
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    93
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    94
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    95
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    96
section {* Records \label{sec:hol-record} *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    97
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    98
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
    99
  In principle, records merely generalize the concept of tuples, where
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   100
  components may be addressed by labels instead of just position.  The
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   101
  logical infrastructure of records in Isabelle/HOL is slightly more
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   102
  advanced, though, supporting truly extensible record schemes.  This
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   103
  admits operations that are polymorphic with respect to record
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   104
  extension, yielding ``object-oriented'' effects like (single)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   105
  inheritance.  See also \cite{NaraschewskiW-TPHOLs98} for more
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   106
  details on object-oriented verification and record subtyping in HOL.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   107
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   108
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   109
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   110
subsection {* Basic concepts *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   111
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   112
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   113
  Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   114
  at the level of terms and types.  The notation is as follows:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   115
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   116
  \begin{center}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   117
  \begin{tabular}{l|l|l}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   118
    & record terms & record types \\ \hline
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   119
    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
   120
    schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   121
      @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   122
  \end{tabular}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   123
  \end{center}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   124
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   125
  \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   126
  "(| x = a |)"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   127
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   128
  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
   129
  @{text a} and field @{text y} of value @{text b}.  The corresponding
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   130
  type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   131
  and @{text "b :: B"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   132
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   133
  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
   134
  @{text x} and @{text y} as before, but also possibly further fields
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   135
  as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   136
  of the syntax).  The improper field ``@{text "\<dots>"}'' of a record
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   137
  scheme is called the \emph{more part}.  Logically it is just a free
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   138
  variable, which is occasionally referred to as ``row variable'' in
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   139
  the literature.  The more part of a record scheme may be
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   140
  instantiated by zero or more further components.  For example, the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   141
  previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z =
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   142
  c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   143
  Fixed records are special instances of record schemes, where
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   144
  ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   145
  element.  In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   146
  for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   147
  
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   148
  \medskip Two key observations make extensible records in a simply
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   149
  typed language like HOL work out:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   150
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   151
  \begin{enumerate}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   152
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   153
  \item the more part is internalized, as a free term or type
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   154
  variable,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   155
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   156
  \item field names are externalized, they cannot be accessed within
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   157
  the logic as first-class values.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   158
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   159
  \end{enumerate}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   160
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   161
  \medskip In Isabelle/HOL record types have to be defined explicitly,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   162
  fixing their field names and types, and their (optional) parent
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   163
  record.  Afterwards, records may be formed using above syntax, while
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   164
  obeying the canonical order of fields as given by their declaration.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   165
  The record package provides several standard operations like
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   166
  selectors and updates.  The common setup for various generic proof
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   167
  tools enable succinct reasoning patterns.  See also the Isabelle/HOL
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   168
  tutorial \cite{isabelle-hol-book} for further instructions on using
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   169
  records in practice.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   170
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   171
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   172
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   173
subsection {* Record specifications *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   174
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   175
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   176
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   177
    @{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   178
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   179
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   180
  \begin{rail}
36158
d2ad76e374d3 HOL record: explicitly allow sort constraints;
wenzelm
parents: 36139
diff changeset
   181
    'record' typespecsorts '=' (type '+')? (constdecl +)
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   182
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   183
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   184
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   185
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   186
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   187
  \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
   188
  \<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
   189
  derived from the optional parent record @{text "\<tau>"} by adding new
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   190
  field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   191
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   192
  The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   193
  covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   194
  \<alpha>\<^sub>m"}.  Type constructor @{text t} has to be new, while @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   195
  \<tau>} needs to specify an instance of an existing record type.  At
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   196
  least one new field @{text "c\<^sub>i"} has to be specified.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   197
  Basically, field names need to belong to a unique record.  This is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   198
  not a real restriction in practice, since fields are qualified by
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   199
  the record name internally.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   200
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   201
  The parent record specification @{text \<tau>} is optional; if omitted
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   202
  @{text t} becomes a root record.  The hierarchy of all records
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   203
  declared within a theory context forms a forest structure, i.e.\ a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   204
  set of trees starting with a root record each.  There is no way to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   205
  merge multiple parent records!
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   206
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   207
  For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   208
  type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   209
  \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   210
  "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   211
  @{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
   212
  \<zeta>\<rparr>"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   213
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   214
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   215
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   216
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   217
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   218
subsection {* Record operations *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   219
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   220
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   221
  Any record definition of the form presented above produces certain
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   222
  standard operations.  Selectors and updates are provided for any
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   223
  field, including the improper one ``@{text more}''.  There are also
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   224
  cumulative record constructor functions.  To simplify the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   225
  presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   226
  \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   227
  \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   228
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   229
  \medskip \textbf{Selectors} and \textbf{updates} are available for
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   230
  any field (including ``@{text more}''):
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   231
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   232
  \begin{matharray}{lll}
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   233
    @{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
   234
    @{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
   235
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   236
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   237
  There is special syntax for application of updates: @{text "r\<lparr>x :=
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   238
  a\<rparr>"} abbreviates term @{text "x_update a r"}.  Further notation for
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   239
  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
   240
  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
   241
  because of postfix notation the order of fields shown here is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   242
  reverse than in the actual term.  Since repeated updates are just
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   243
  function applications, fields may be freely permuted in @{text "\<lparr>x
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   244
  := a, y := b, z := c\<rparr>"}, as far as logical equality is concerned.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   245
  Thus commutativity of independent updates can be proven within the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   246
  logic for any two fields, but not as a general theorem.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   247
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   248
  \medskip The \textbf{make} operation provides a cumulative record
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   249
  constructor function:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   250
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   251
  \begin{matharray}{lll}
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   252
    @{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
   253
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   254
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   255
  \medskip We now reconsider the case of non-root records, which are
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   256
  derived of some parent.  In general, the latter may depend on
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   257
  another parent as well, resulting in a list of \emph{ancestor
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   258
  records}.  Appending the lists of fields of all ancestors results in
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   259
  a certain field prefix.  The record package automatically takes care
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   260
  of this by lifting operations over this context of ancestor fields.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   261
  Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   262
  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
   263
  the above record operations will get the following types:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   264
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   265
  \medskip
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   266
  \begin{tabular}{lll}
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   267
    @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   268
    @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow> 
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   269
      \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   270
      \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   271
    @{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
   272
      \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   273
  \end{tabular}
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   274
  \medskip
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   275
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   276
  \noindent Some further operations address the extension aspect of a
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   277
  derived record scheme specifically: @{text "t.fields"} produces a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   278
  record fragment consisting of exactly the new fields introduced here
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   279
  (the result may serve as a more part elsewhere); @{text "t.extend"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   280
  takes a fixed record and adds a given more part; @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   281
  "t.truncate"} restricts a record scheme to a fixed record.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   282
26852
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   283
  \medskip
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   284
  \begin{tabular}{lll}
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   285
    @{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
   286
    @{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   287
      \<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   288
    @{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
   289
  \end{tabular}
a31203f58b20 misc tuning;
wenzelm
parents: 26849
diff changeset
   290
  \medskip
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   291
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   292
  \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   293
  for root records.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   294
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   295
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   296
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   297
subsection {* Derived rules and proof tools *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   298
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   299
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   300
  The record package proves several results internally, declaring
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   301
  these facts to appropriate proof tools.  This enables users to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   302
  reason about record structures quite conveniently.  Assume that
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   303
  @{text t} is a record type as specified above.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   304
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   305
  \begin{enumerate}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   306
  
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   307
  \item Standard conversions for selectors or updates applied to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   308
  record constructor terms are made part of the default Simplifier
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   309
  context; thus proofs by reduction of basic operations merely require
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   310
  the @{method simp} method without further arguments.  These rules
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   311
  are available as @{text "t.simps"}, too.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   312
  
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   313
  \item Selectors applied to updated records are automatically reduced
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   314
  by an internal simplification procedure, which is also part of the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   315
  standard Simplifier setup.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   316
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   317
  \item Inject equations of a form analogous to @{prop "(x, y) = (x',
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   318
  y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   319
  Reasoner as @{attribute iff} rules.  These rules are available as
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   320
  @{text "t.iffs"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   321
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   322
  \item The introduction rule for record equality analogous to @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   323
  "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
   324
  and as the basic rule context as ``@{attribute intro}@{text "?"}''.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   325
  The rule is called @{text "t.equality"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   326
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   327
  \item Representations of arbitrary record expressions as canonical
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   328
  constructor terms are provided both in @{method cases} and @{method
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   329
  induct} format (cf.\ the generic proof methods of the same name,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   330
  \secref{sec:cases-induct}).  Several variations are available, for
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   331
  fixed records, record schemes, more parts etc.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   332
  
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   333
  The generic proof methods are sufficiently smart to pick the most
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   334
  sensible rule according to the type of the indicated record
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   335
  expression: users just need to apply something like ``@{text "(cases
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   336
  r)"}'' to a certain proof problem.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   337
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   338
  \item The derived record operations @{text "t.make"}, @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   339
  "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   340
  treated automatically, but usually need to be expanded by hand,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   341
  using the collective fact @{text "t.defs"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   342
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   343
  \end{enumerate}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   344
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   345
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   346
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   347
section {* Datatypes \label{sec:hol-datatype} *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   348
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   349
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   350
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   351
    @{command_def (HOL) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   352
  @{command_def (HOL) "rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   353
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   354
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   355
  \begin{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   356
    'datatype' (dtspec + 'and')
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   357
    ;
27452
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
   358
    'rep\_datatype' ('(' (name +) ')')? (term +)
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   359
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   360
35351
7425aece4ee3 allow general mixfix syntax for type constructors;
wenzelm
parents: 35331
diff changeset
   361
    dtspec: parname? typespec mixfix? '=' (cons + '|')
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   362
    ;
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   363
    cons: name ( type * ) mixfix?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   364
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   365
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   366
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   367
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   368
  \item @{command (HOL) "datatype"} defines inductive datatypes in
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   369
  HOL.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   370
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   371
  \item @{command (HOL) "rep_datatype"} represents existing types as
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   372
  inductive ones, generating the standard infrastructure of derived
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   373
  concepts (primitive recursion etc.).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   374
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   375
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   376
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   377
  The induction and exhaustion theorems generated provide case names
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   378
  according to the constructors involved, while parameters are named
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   379
  after the types (see also \secref{sec:cases-induct}).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   380
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   381
  See \cite{isabelle-HOL} for more details on datatypes, but beware of
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   382
  the old-style theory syntax being used there!  Apart from proper
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   383
  proof methods for case-analysis and induction, there are also
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   384
  emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   385
  induct_tac} available, see \secref{sec:hol-induct-tac}; these admit
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   386
  to refer directly to the internal structure of subgoals (including
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   387
  internally bound parameters).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   388
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   389
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   390
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   391
section {* Recursive functions \label{sec:recursion} *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   392
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   393
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   394
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   395
    @{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
   396
    @{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
   397
    @{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
   398
    @{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   399
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   400
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   401
  \begin{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   402
    'primrec' target? fixes 'where' equations
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   403
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   404
    equations: (thmdecl? prop + '|')
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   405
    ;
26985
51c5acd57b75 function: uniform treatment of target, not as config;
wenzelm
parents: 26894
diff changeset
   406
    ('fun' | 'function') target? functionopts? fixes 'where' clauses
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   407
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   408
    clauses: (thmdecl? prop ('(' 'otherwise' ')')? + '|')
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   409
    ;
26985
51c5acd57b75 function: uniform treatment of target, not as config;
wenzelm
parents: 26894
diff changeset
   410
    functionopts: '(' (('sequential' | 'domintros' | 'tailrec' | 'default' term) + ',') ')'
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   411
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   412
    'termination' ( term )?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   413
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   414
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   415
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   416
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   417
  \item @{command (HOL) "primrec"} defines primitive recursive
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   418
  functions over datatypes, see also \cite{isabelle-HOL}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   419
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   420
  \item @{command (HOL) "function"} defines functions by general
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   421
  wellfounded recursion. A detailed description with examples can be
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   422
  found in \cite{isabelle-function}. The function is specified by a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   423
  set of (possibly conditional) recursive equations with arbitrary
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   424
  pattern matching. The command generates proof obligations for the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   425
  completeness and the compatibility of patterns.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   426
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   427
  The defined function is considered partial, and the resulting
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   428
  simplification rules (named @{text "f.psimps"}) and induction rule
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   429
  (named @{text "f.pinduct"}) are guarded by a generated domain
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   430
  predicate @{text "f_dom"}. The @{command (HOL) "termination"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   431
  command can then be used to establish that the function is total.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   432
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   433
  \item @{command (HOL) "fun"} is a shorthand notation for ``@{command
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   434
  (HOL) "function"}~@{text "(sequential)"}, followed by automated
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   435
  proof attempts regarding pattern matching and termination.  See
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   436
  \cite{isabelle-function} for further details.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   437
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   438
  \item @{command (HOL) "termination"}~@{text f} commences a
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   439
  termination proof for the previously defined function @{text f}.  If
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   440
  this is omitted, the command refers to the most recent function
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   441
  definition.  After the proof is closed, the recursive equations and
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   442
  the induction principle is established.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   443
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   444
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   445
27452
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
   446
  Recursive definitions introduced by the @{command (HOL) "function"}
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
   447
  command accommodate
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   448
  reasoning by induction (cf.\ \secref{sec:cases-induct}): rule @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   449
  "c.induct"} (where @{text c} is the name of the function definition)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   450
  refers to a specific induction rule, with parameters named according
33857
0cb5002c52db clarified; checked
krauss
parents: 31998
diff changeset
   451
  to the user-specified equations. Cases are numbered (starting from 1).
0cb5002c52db clarified; checked
krauss
parents: 31998
diff changeset
   452
0cb5002c52db clarified; checked
krauss
parents: 31998
diff changeset
   453
  For @{command (HOL) "primrec"}, the induction principle coincides
27452
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
   454
  with structural recursion on the datatype the recursion is carried
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
   455
  out.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   456
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   457
  The equations provided by these packages may be referred later as
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   458
  theorem list @{text "f.simps"}, where @{text f} is the (collective)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   459
  name of the functions defined.  Individual equations may be named
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   460
  explicitly as well.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   461
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   462
  The @{command (HOL) "function"} command accepts the following
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   463
  options.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   464
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   465
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   466
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   467
  \item @{text sequential} enables a preprocessor which disambiguates
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   468
  overlapping patterns by making them mutually disjoint.  Earlier
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   469
  equations take precedence over later ones.  This allows to give the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   470
  specification in a format very similar to functional programming.
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   471
  Note that the resulting simplification and induction rules
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   472
  correspond to the transformed specification, not the one given
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   473
  originally. This usually means that each equation given by the user
36139
0c2538afe8e8 Spelling error: theroems -> theorems
hoelzl
parents: 35841
diff changeset
   474
  may result in several theorems.  Also note that this automatic
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   475
  transformation only works for ML-style datatype patterns.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   476
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   477
  \item @{text domintros} enables the automated generation of
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   478
  introduction rules for the domain predicate. While mostly not
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   479
  needed, they can be helpful in some proofs about partial functions.
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
  \item @{text tailrec} generates the unconstrained recursive
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   482
  equations even without a termination proof, provided that the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   483
  function is tail-recursive. This currently only works
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   484
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   485
  \item @{text "default d"} allows to specify a default value for a
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   486
  (partial) function, which will ensure that @{text "f x = d x"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   487
  whenever @{text "x \<notin> f_dom"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   488
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   489
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   490
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   491
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   492
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   493
subsection {* Proof methods related to recursive definitions *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   494
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   495
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   496
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   497
    @{method_def (HOL) pat_completeness} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   498
    @{method_def (HOL) relation} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   499
    @{method_def (HOL) lexicographic_order} & : & @{text method} \\
33858
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   500
    @{method_def (HOL) size_change} & : & @{text method} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   501
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   502
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   503
  \begin{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   504
    'relation' term
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   505
    ;
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   506
    'lexicographic\_order' ( clasimpmod * )
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   507
    ;
33858
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   508
    'size\_change' ( orders ( clasimpmod * ) )
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   509
    ;
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   510
    orders: ( 'max' | 'min' | 'ms' ) *
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   511
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   512
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   513
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   514
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   515
  \item @{method (HOL) pat_completeness} is a specialized method to
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   516
  solve goals regarding the completeness of pattern matching, as
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   517
  required by the @{command (HOL) "function"} package (cf.\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   518
  \cite{isabelle-function}).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   519
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   520
  \item @{method (HOL) relation}~@{text R} introduces a termination
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   521
  proof using the relation @{text R}.  The resulting proof state will
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   522
  contain goals expressing that @{text R} is wellfounded, and that the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   523
  arguments of recursive calls decrease with respect to @{text R}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   524
  Usually, this method is used as the initial proof step of manual
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   525
  termination proofs.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   526
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   527
  \item @{method (HOL) "lexicographic_order"} attempts a fully
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   528
  automated termination proof by searching for a lexicographic
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   529
  combination of size measures on the arguments of the function. The
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   530
  method accepts the same arguments as the @{method auto} method,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   531
  which it uses internally to prove local descents.  The same context
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   532
  modifiers as for @{method auto} are accepted, see
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   533
  \secref{sec:clasimp}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   534
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   535
  In case of failure, extensive information is printed, which can help
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   536
  to analyse the situation (cf.\ \cite{isabelle-function}).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   537
33858
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   538
  \item @{method (HOL) "size_change"} also works on termination goals,
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   539
  using a variation of the size-change principle, together with a
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   540
  graph decomposition technique (see \cite{krauss_phd} for details).
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   541
  Three kinds of orders are used internally: @{text max}, @{text min},
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   542
  and @{text ms} (multiset), which is only available when the theory
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   543
  @{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
   544
  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
   545
  internally.
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   546
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   547
 For local descent proofs, the same context modifiers as for @{method
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   548
  auto} are accepted, see \secref{sec:clasimp}.
0c348f7997f7 documented size_change in isar-ref manual
krauss
parents: 33857
diff changeset
   549
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   550
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   551
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   552
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   553
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   554
subsection {* Old-style recursive function definitions (TFL) *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   555
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   556
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   557
  The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   558
  "recdef_tc"} for defining recursive are mostly obsolete; @{command
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   559
  (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   560
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   561
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   562
    @{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
   563
    @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   564
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   565
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   566
  \begin{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   567
    'recdef' ('(' 'permissive' ')')? \\ name term (prop +) hints?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   568
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   569
    recdeftc thmdecl? tc
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   570
    ;
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   571
    hints: '(' 'hints' ( recdefmod * ) ')'
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   572
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   573
    recdefmod: (('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   574
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   575
    tc: nameref ('(' nat ')')?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   576
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   577
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   578
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   579
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   580
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   581
  \item @{command (HOL) "recdef"} defines general well-founded
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   582
  recursive functions (using the TFL package), see also
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   583
  \cite{isabelle-HOL}.  The ``@{text "(permissive)"}'' option tells
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   584
  TFL to recover from failed proof attempts, returning unfinished
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   585
  results.  The @{text recdef_simp}, @{text recdef_cong}, and @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   586
  recdef_wf} hints refer to auxiliary rules to be used in the internal
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   587
  automated proof process of TFL.  Additional @{syntax clasimpmod}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   588
  declarations (cf.\ \secref{sec:clasimp}) may be given to tune the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   589
  context of the Simplifier (cf.\ \secref{sec:simplifier}) and
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   590
  Classical reasoner (cf.\ \secref{sec:classical}).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   591
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   592
  \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   593
  proof for leftover termination condition number @{text i} (default
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   594
  1) as generated by a @{command (HOL) "recdef"} definition of
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   595
  constant @{text c}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   596
  
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   597
  Note that in most cases, @{command (HOL) "recdef"} is able to finish
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   598
  its internal proofs without manual intervention.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   599
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   600
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   601
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   602
  \medskip Hints for @{command (HOL) "recdef"} may be also declared
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   603
  globally, using the following attributes.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   604
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   605
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   606
    @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   607
    @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   608
    @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   609
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   610
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   611
  \begin{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   612
    ('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del')
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   613
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   614
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   615
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   616
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   617
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   618
section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   619
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   620
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   621
  An \textbf{inductive definition} specifies the least predicate (or
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   622
  set) @{text R} closed under given rules: applying a rule to elements
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   623
  of @{text R} yields a result within @{text R}.  For example, a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   624
  structural operational semantics is an inductive definition of an
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   625
  evaluation relation.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   626
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   627
  Dually, a \textbf{coinductive definition} specifies the greatest
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   628
  predicate~/ set @{text R} that is consistent with given rules: every
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   629
  element of @{text R} can be seen as arising by applying a rule to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   630
  elements of @{text R}.  An important example is using bisimulation
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   631
  relations to formalise equivalence of processes and infinite data
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   632
  structures.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   633
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   634
  \medskip The HOL package is related to the ZF one, which is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   635
  described in a separate paper,\footnote{It appeared in CADE
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   636
  \cite{paulson-CADE}; a longer version is distributed with Isabelle.}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   637
  which you should refer to in case of difficulties.  The package is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   638
  simpler than that of ZF thanks to implicit type-checking in HOL.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   639
  The types of the (co)inductive predicates (or sets) determine the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   640
  domain of the fixedpoint definition, and the package does not have
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   641
  to use inference rules for type-checking.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   642
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   643
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   644
    @{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
   645
    @{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
   646
    @{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
   647
    @{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
   648
    @{attribute_def (HOL) mono} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   649
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   650
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   651
  \begin{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   652
    ('inductive' | 'inductive\_set' | 'coinductive' | 'coinductive\_set') target? fixes ('for' fixes)? \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   653
    ('where' clauses)? ('monos' thmrefs)?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   654
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   655
    clauses: (thmdecl? prop + '|')
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   656
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   657
    'mono' (() | 'add' | 'del')
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   658
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   659
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   660
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   661
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   662
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   663
  \item @{command (HOL) "inductive"} and @{command (HOL)
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   664
  "coinductive"} define (co)inductive predicates from the
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   665
  introduction rules given in the @{keyword "where"} part.  The
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   666
  optional @{keyword "for"} part contains a list of parameters of the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   667
  (co)inductive predicates that remain fixed throughout the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   668
  definition.  The optional @{keyword "monos"} section contains
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   669
  \emph{monotonicity theorems}, which are required for each operator
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   670
  applied to a recursive set in the introduction rules.  There
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   671
  \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
   672
  for each premise @{text "M R\<^sub>i t"} in an introduction rule!
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   673
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   674
  \item @{command (HOL) "inductive_set"} and @{command (HOL)
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   675
  "coinductive_set"} are wrappers for to the previous commands,
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   676
  allowing the definition of (co)inductive sets.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   677
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   678
  \item @{attribute (HOL) mono} declares monotonicity rules.  These
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   679
  rule are involved in the automated monotonicity proof of @{command
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   680
  (HOL) "inductive"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   681
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   682
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   683
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   684
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   685
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   686
subsection {* Derived rules *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   687
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   688
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   689
  Each (co)inductive definition @{text R} adds definitions to the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   690
  theory and also proves some theorems:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   691
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   692
  \begin{description}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   693
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   694
  \item @{text R.intros} is the list of introduction rules as proven
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   695
  theorems, for the recursive predicates (or sets).  The rules are
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   696
  also available individually, using the names given them in the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   697
  theory file;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   698
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   699
  \item @{text R.cases} is the case analysis (or elimination) rule;
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   700
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   701
  \item @{text R.induct} or @{text R.coinduct} is the (co)induction
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   702
  rule.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   703
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   704
  \end{description}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   705
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   706
  When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   707
  defined simultaneously, the list of introduction rules is called
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   708
  @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   709
  called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   710
  of mutual induction rules is called @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   711
  "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
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
subsection {* Monotonicity theorems *}
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
  Each theory contains a default set of theorems that are used in
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   719
  monotonicity proofs.  New rules can be added to this set via the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   720
  @{attribute (HOL) mono} attribute.  The HOL theory @{text Inductive}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   721
  shows how this is done.  In general, the following monotonicity
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   722
  theorems may be added:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   723
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   724
  \begin{itemize}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   725
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   726
  \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
   727
  monotonicity of inductive definitions whose introduction rules have
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   728
  premises involving terms such as @{text "M R\<^sub>i t"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   729
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   730
  \item Monotonicity theorems for logical operators, which are of the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   731
  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
   732
  the case of the operator @{text "\<or>"}, the corresponding theorem is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   733
  \[
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   734
  \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
   735
  \]
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   736
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   737
  \item De Morgan style equations for reasoning about the ``polarity''
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   738
  of expressions, e.g.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   739
  \[
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   740
  @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   741
  @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   742
  \]
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   743
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   744
  \item Equations for reducing complex operators to more primitive
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   745
  ones whose monotonicity can easily be proved, e.g.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   746
  \[
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   747
  @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   748
  @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   749
  \]
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   750
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   751
  \end{itemize}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   752
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   753
  %FIXME: Example of an inductive definition
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   754
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   755
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   756
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   757
section {* Arithmetic proof support *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   758
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   759
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   760
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   761
    @{method_def (HOL) arith} & : & @{text method} \\
30863
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   762
    @{attribute_def (HOL) arith} & : & @{text attribute} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   763
    @{attribute_def (HOL) arith_split} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   764
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   765
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   766
  The @{method (HOL) arith} method decides linear arithmetic problems
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   767
  (on types @{text nat}, @{text int}, @{text real}).  Any current
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   768
  facts are inserted into the goal before running the procedure.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   769
30863
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   770
  The @{attribute (HOL) arith} attribute declares facts that are
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   771
  always supplied to the arithmetic provers implicitly.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   772
30863
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   773
  The @{attribute (HOL) arith_split} attribute declares case split
30865
5106e13d400f fixed formal markup;
wenzelm
parents: 30863
diff changeset
   774
  rules to be expanded before @{method (HOL) arith} is invoked.
30863
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   775
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   776
  Note that a simpler (but faster) arithmetic prover is
5dc392a59bb7 Finite_Set: lemma
nipkow
parents: 30242
diff changeset
   777
  already invoked by the Simplifier.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   778
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   779
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   780
30169
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   781
section {* Intuitionistic proof search *}
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   782
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   783
text {*
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   784
  \begin{matharray}{rcl}
30171
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   785
    @{method_def (HOL) iprover} & : & @{text method} \\
30169
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   786
  \end{matharray}
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   787
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   788
  \begin{rail}
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 35351
diff changeset
   789
    'iprover' ( rulemod * )
30169
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   790
    ;
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   791
  \end{rail}
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   792
30171
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   793
  The @{method (HOL) iprover} method performs intuitionistic proof
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   794
  search, depending on specifically declared rules from the context,
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   795
  or given as explicit arguments.  Chained facts are inserted into the
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 35351
diff changeset
   796
  goal before commencing proof search.
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 35351
diff changeset
   797
30169
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   798
  Rules need to be classified as @{attribute (Pure) intro},
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   799
  @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   800
  ``@{text "!"}'' indicator refers to ``safe'' rules, which may be
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   801
  applied aggressively (without considering back-tracking later).
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   802
  Rules declared with ``@{text "?"}'' are ignored in proof search (the
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   803
  single-step @{method rule} method still observes these).  An
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   804
  explicit weight annotation may be given as well; otherwise the
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   805
  number of rule premises will be taken into account here.
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   806
*}
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   807
9531eaafd781 moved method "iprover" to HOL specific part;
wenzelm
parents: 29560
diff changeset
   808
30171
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   809
section {* Coherent Logic *}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   810
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   811
text {*
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   812
  \begin{matharray}{rcl}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   813
    @{method_def (HOL) "coherent"} & : & @{text method} \\
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   814
  \end{matharray}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   815
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   816
  \begin{rail}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   817
    'coherent' thmrefs?
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   818
    ;
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   819
  \end{rail}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   820
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   821
  The @{method (HOL) coherent} method solves problems of
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   822
  \emph{Coherent Logic} \cite{Bezem-Coquand:2005}, which covers
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   823
  applications in confluence theory, lattice theory and projective
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   824
  geometry.  See @{"file" "~~/src/HOL/ex/Coherent.thy"} for some
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   825
  examples.
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   826
*}
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   827
5989863ffafc added method "coherent";
wenzelm
parents: 30169
diff changeset
   828
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   829
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
   830
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   831
text {*
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   832
  Identifying incorrect propositions usually involves evaluation of
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   833
  particular assignments and systematic counter example search.  This
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   834
  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
   835
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   836
  \begin{matharray}{rcl}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   837
    @{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
   838
    @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   839
    @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \<rightarrow> theory"}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   840
  \end{matharray}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   841
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   842
  \begin{rail}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   843
    'value' ( ( '[' name ']' ) ? ) modes? term
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   844
    ;
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   845
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   846
    'quickcheck' ( ( '[' args ']' ) ? ) nat?
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   847
    ;
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   848
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   849
    'quickcheck_params' ( ( '[' args ']' ) ? )
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   850
    ;
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   851
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   852
    modes: '(' (name + ) ')'
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   853
    ;
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   854
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   855
    args: ( name '=' value + ',' )
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   856
    ;
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   857
  \end{rail}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   858
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   859
  \begin{description}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   860
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   861
  \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
   862
    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
   863
    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
   864
    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
   865
    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
   866
    Alternatively a specific evaluator can be selected using square
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   867
    brackets; available evaluators include @{text nbe} for
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   868
    \emph{normalization by evaluation} and \emph{code} for code
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   869
    generation in SML.
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   870
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   871
  \item @{command (HOL) "quickcheck"} tests the current goal for
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   872
    counter examples using a series of arbitrary assignments for its
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   873
    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
   874
    can be selected explicitly using an optional goal index.
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   875
    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
   876
    @{command (HOL) "quickcheck"}, notably:
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   877
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   878
    \begin{description}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   879
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   880
      \item[size] specifies the maximum size of the search space for
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   881
        assignment values.
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   882
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   883
      \item[iterations] sets how many sets of assignments are
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   884
        generated for each particular size.
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   885
35331
450ab945c451 document Quickcheck's "no_assms" option
blanchet
parents: 34172
diff changeset
   886
      \item[no\_assms] specifies whether assumptions in
450ab945c451 document Quickcheck's "no_assms" option
blanchet
parents: 34172
diff changeset
   887
        structured proofs should be ignored.
450ab945c451 document Quickcheck's "no_assms" option
blanchet
parents: 34172
diff changeset
   888
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   889
    \end{description}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   890
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   891
    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
   892
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   893
  \item @{command (HOL) "quickcheck_params"} changes quickcheck
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   894
    configuration options persitently.
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   895
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   896
  \end{description}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   897
*}
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   898
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   899
35744
93603d7b8ee9 removed obsolete HOL 'typedecl';
wenzelm
parents: 35613
diff changeset
   900
section {* Invoking automated reasoning tools --- The Sledgehammer *}
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   901
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   902
text {*
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   903
  Isabelle/HOL includes a generic \emph{ATP manager} that allows
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   904
  external automated reasoning tools to crunch a pending goal.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   905
  Supported provers include E\footnote{\url{http://www.eprover.org}},
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   906
  SPASS\footnote{\url{http://www.spass-prover.org/}}, and Vampire.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   907
  There is also a wrapper to invoke provers remotely via the
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   908
  SystemOnTPTP\footnote{\url{http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTP}}
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   909
  web service.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   910
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   911
  The problem passed to external provers consists of the goal together
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   912
  with a smart selection of lemmas from the current theory context.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   913
  The result of a successful proof search is some source text that
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   914
  usually reconstructs the proof within Isabelle, without requiring
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   915
  external provers again.  The Metis
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   916
  prover\footnote{\url{http://www.gilith.com/software/metis/}} that is
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   917
  integrated into Isabelle/HOL is being used here.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   918
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   919
  In this mode of operation, heavy means of automated reasoning are
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   920
  used as a strong relevance filter, while the main proof checking
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   921
  works via explicit inferences going through the Isabelle kernel.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   922
  Moreover, rechecking Isabelle proof texts with already specified
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   923
  auxiliary facts is much faster than performing fully automated
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   924
  search over and over again.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   925
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   926
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   927
    @{command_def (HOL) "sledgehammer"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   928
    @{command_def (HOL) "print_atps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   929
    @{command_def (HOL) "atp_info"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   930
    @{command_def (HOL) "atp_kill"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
29112
f2b45eea6dac added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents: 28761
diff changeset
   931
    @{command_def (HOL) "atp_messages"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   932
    @{method_def (HOL) metis} & : & @{text method} \\
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   933
  \end{matharray}
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   934
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   935
  \begin{rail}
31912
f5bd306f5e9d more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents: 31254
diff changeset
   936
  'sledgehammer' ( nameref * )
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   937
  ;
29112
f2b45eea6dac added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents: 28761
diff changeset
   938
  'atp\_messages' ('(' nat ')')?
29114
715178f6ae31 repaired railroad accident;
wenzelm
parents: 29112
diff changeset
   939
  ;
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   940
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   941
  'metis' thmrefs
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   942
  ;
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   943
  \end{rail}
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   944
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   945
  \begin{description}
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   946
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   947
  \item @{command (HOL) sledgehammer}~@{text "prover\<^sub>1 \<dots> prover\<^sub>n"}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   948
  invokes the specified automated theorem provers on the first
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   949
  subgoal.  Provers are run in parallel, the first successful result
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   950
  is displayed, and the other attempts are terminated.
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   951
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   952
  Provers are defined in the theory context, see also @{command (HOL)
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   953
  print_atps}.  If no provers are given as arguments to @{command
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   954
  (HOL) sledgehammer}, the system refers to the default defined as
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   955
  ``ATP provers'' preference by the user interface.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   956
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   957
  There are additional preferences for timeout (default: 60 seconds),
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   958
  and the maximum number of independent prover processes (default: 5);
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   959
  excessive provers are automatically terminated.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   960
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   961
  \item @{command (HOL) print_atps} prints the list of automated
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   962
  theorem provers available to the @{command (HOL) sledgehammer}
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   963
  command.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   964
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   965
  \item @{command (HOL) atp_info} prints information about presently
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   966
  running provers, including elapsed runtime, and the remaining time
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   967
  until timeout.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   968
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   969
  \item @{command (HOL) atp_kill} terminates all presently running
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   970
  provers.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   971
29112
f2b45eea6dac added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents: 28761
diff changeset
   972
  \item @{command (HOL) atp_messages} displays recent messages issued
f2b45eea6dac added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents: 28761
diff changeset
   973
  by automated theorem provers.  This allows to examine results that
f2b45eea6dac added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents: 28761
diff changeset
   974
  might have got lost due to the asynchronous nature of default
f2b45eea6dac added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents: 28761
diff changeset
   975
  @{command (HOL) sledgehammer} output.  An optional message limit may
f2b45eea6dac added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents: 28761
diff changeset
   976
  be specified (default 5).
f2b45eea6dac added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents: 28761
diff changeset
   977
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   978
  \item @{method (HOL) metis}~@{text "facts"} invokes the Metis prover
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   979
  with the given facts.  Metis is an automated proof tool of medium
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   980
  strength, but is fully integrated into Isabelle/HOL, with explicit
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   981
  inferences going through the kernel.  Thus its results are
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   982
  guaranteed to be ``correct by construction''.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   983
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   984
  Note that all facts used with Metis need to be specified as explicit
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   985
  arguments.  There are no rule declarations as for other Isabelle
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   986
  provers, like @{method blast} or @{method fast}.
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   987
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
   988
  \end{description}
28603
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   989
*}
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   990
b40800eef8a7 added sledgehammer etc.;
wenzelm
parents: 28562
diff changeset
   991
28752
wenzelm
parents: 28687
diff changeset
   992
section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   993
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   994
text {*
27123
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
   995
  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
   996
  induction in unstructured tactic scripts; see also
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
   997
  \secref{sec:cases-induct} for proper Isar versions of similar ideas.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   998
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
   999
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1000
    @{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
  1001
    @{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
  1002
    @{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
  1003
    @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1004
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1005
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1006
  \begin{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1007
    'case\_tac' goalspec? term rule?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1008
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1009
    'induct\_tac' goalspec? (insts * 'and') rule?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1010
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1011
    'ind\_cases' (prop +) ('for' (name +)) ?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1012
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1013
    'inductive\_cases' (thmdecl? (prop +) + 'and')
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1014
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1015
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1016
    rule: ('rule' ':' thmref)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1017
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1018
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1019
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1020
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1021
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1022
  \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1023
  to reason about inductive types.  Rules are selected according to
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1024
  the declarations by the @{attribute cases} and @{attribute induct}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1025
  attributes, cf.\ \secref{sec:cases-induct}.  The @{command (HOL)
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1026
  datatype} package already takes care of this.
27123
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1027
11fcdd5897dd case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents: 27103
diff changeset
  1028
  These unstructured tactics feature both goal addressing and dynamic
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1029
  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
  1030
  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
  1031
  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
  1032
  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
  1033
  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
  1034
  being addressed.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1035
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1036
  \item @{method (HOL) ind_cases} and @{command (HOL)
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1037
  "inductive_cases"} provide an interface to the internal @{ML_text
26860
7c749112261c replaced some latex macros by antiquotations;
wenzelm
parents: 26852
diff changeset
  1038
  mk_cases} operation.  Rules are simplified in an unrestricted
7c749112261c replaced some latex macros by antiquotations;
wenzelm
parents: 26852
diff changeset
  1039
  forward manner.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1040
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1041
  While @{method (HOL) ind_cases} is a proof method to apply the
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1042
  result immediately as elimination rules, @{command (HOL)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1043
  "inductive_cases"} provides case split theorems at the theory level
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1044
  for later use.  The @{keyword "for"} argument of the @{method (HOL)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1045
  ind_cases} method allows to specify a list of variables that should
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1046
  be generalized before applying the resulting rule.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1047
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1048
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1049
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1050
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1051
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1052
section {* Executable code *}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1053
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1054
text {*
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1055
  Isabelle/Pure provides two generic frameworks to support code
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1056
  generation from executable specifications.  Isabelle/HOL
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1057
  instantiates these mechanisms in a way that is amenable to end-user
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1058
  applications.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1059
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1060
  One framework generates code from both functional and relational
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1061
  programs to SML.  See \cite{isabelle-HOL} for further information
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1062
  (this actually covers the new-style theory format as well).
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1063
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1064
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1065
    @{command_def (HOL) "code_module"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1066
    @{command_def (HOL) "code_library"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1067
    @{command_def (HOL) "consts_code"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1068
    @{command_def (HOL) "types_code"} & : & @{text "theory \<rightarrow> theory"} \\  
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1069
    @{attribute_def (HOL) code} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1070
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1071
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1072
  \begin{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1073
  ( 'code\_module' | 'code\_library' ) modespec ? name ? \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1074
    ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1075
    'contains' ( ( name '=' term ) + | term + )
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1076
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1077
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1078
  modespec: '(' ( name * ) ')'
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1079
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1080
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1081
  'consts\_code' (codespec +)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1082
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1083
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1084
  codespec: const template attachment ?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1085
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1086
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1087
  'types\_code' (tycodespec +)
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1088
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1089
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1090
  tycodespec: name template attachment ?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1091
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1092
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1093
  const: term
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1094
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1095
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1096
  template: '(' string ')'
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1097
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1098
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1099
  attachment: 'attach' modespec ? verblbrace text verbrbrace
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1100
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1101
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1102
  'code' (name)?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1103
  ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1104
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1105
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1106
  \medskip The other framework generates code from functional programs
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1107
  (including overloading using type classes) to SML \cite{SML}, OCaml
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1108
  \cite{OCaml} and Haskell \cite{haskell-revised-report}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1109
  Conceptually, code generation is split up in three steps:
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1110
  \emph{selection} of code theorems, \emph{translation} into an
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1111
  abstract executable view and \emph{serialization} to a specific
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1112
  \emph{target language}.  See \cite{isabelle-codegen} for an
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1113
  introduction on how to use it.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1114
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1115
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1116
    @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1117
    @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1118
    @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1119
    @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1120
    @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1121
    @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1122
    @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1123
    @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1124
    @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1125
    @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1126
    @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1127
    @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1128
    @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1129
    @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
31254
03a35fbc9dc6 documented print_codeproc command
haftmann
parents: 30865
diff changeset
  1130
    @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1131
    @{attribute_def (HOL) code} & : & @{text attribute} \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1132
  \end{matharray}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1133
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1134
  \begin{rail}
34172
4301e9ea1c54 updated documentation
haftmann
parents: 33858
diff changeset
  1135
    'export\_code' ( constexpr + ) \\
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1136
      ( ( 'in' target ( 'module\_name' string ) ? \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1137
        ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1138
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1139
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1140
    'code\_thms' ( constexpr + ) ?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1141
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1142
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1143
    'code\_deps' ( constexpr + ) ?
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1144
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1145
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1146
    const: term
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1147
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1148
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1149
    constexpr: ( const | 'name.*' | '*' )
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1150
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1151
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1152
    typeconstructor: nameref
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1153
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1154
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1155
    class: nameref
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1156
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1157
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1158
    target: 'OCaml' | 'SML' | 'Haskell'
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1159
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1160
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1161
    'code\_datatype' const +
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1162
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1163
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1164
    'code\_const' (const + 'and') \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1165
      ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1166
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1167
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1168
    'code\_type' (typeconstructor + 'and') \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1169
      ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1170
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1171
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1172
    'code\_class' (class + 'and') \\
28687
150a8a1eae1a simplified syntax for class parameters
haftmann
parents: 28603
diff changeset
  1173
      ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1174
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1175
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1176
    'code\_instance' (( typeconstructor '::' class ) + 'and') \\
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1177
      ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1178
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1179
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1180
    'code\_monad' const const target
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1181
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1182
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1183
    'code\_reserved' target ( string + )
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
    'code\_include' target ( string ( string | '-') )
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1187
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1188
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1189
    'code\_modulename' target ( ( string string ) + )
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1190
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1191
27452
5c1fb7d262bf adjusted rep_datatype
haftmann
parents: 27123
diff changeset
  1192
    'code\_abort' ( const + )
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1193
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1194
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1195
    syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1196
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1197
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1198
    'code' ( 'del' ) ?
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1199
    ;
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1200
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1201
    'code_unfold' ( 'del' ) ?
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1202
    ;
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1203
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1204
    'code_post' ( 'del' ) ?
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1205
    ;
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1206
  \end{rail}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1207
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1208
  \begin{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1209
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1210
  \item @{command (HOL) "export_code"} is the canonical interface for
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1211
  generating and serializing code: for a given list of constants, code
34172
4301e9ea1c54 updated documentation
haftmann
parents: 33858
diff changeset
  1212
  is generated for the specified target languages.  If no serialization
4301e9ea1c54 updated documentation
haftmann
parents: 33858
diff changeset
  1213
  instruction is given, only abstract code is generated internally.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1214
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1215
  Constants may be specified by giving them literally, referring to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1216
  all executable contants within a certain theory by giving @{text
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1217
  "name.*"}, or referring to \emph{all} executable constants currently
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1218
  available by giving @{text "*"}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1219
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1220
  By default, for each involved theory one corresponding name space
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1221
  module is generated.  Alternativly, a module name may be specified
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1222
  after the @{keyword "module_name"} keyword; then \emph{all} code is
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1223
  placed in this module.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1224
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1225
  For \emph{SML} and \emph{OCaml}, the file specification refers to a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1226
  single file; for \emph{Haskell}, it refers to a whole directory,
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1227
  where code is generated in multiple files reflecting the module
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1228
  hierarchy.  The file specification ``@{text "-"}'' denotes standard
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1229
  output.  For \emph{SML}, omitting the file specification compiles
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1230
  code internally in the context of the current ML session.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1231
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1232
  Serializers take an optional list of arguments in parentheses.  For
34172
4301e9ea1c54 updated documentation
haftmann
parents: 33858
diff changeset
  1233
  \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits
4301e9ea1c54 updated documentation
haftmann
parents: 33858
diff changeset
  1234
  explicit module signatures.
4301e9ea1c54 updated documentation
haftmann
parents: 33858
diff changeset
  1235
  
4301e9ea1c54 updated documentation
haftmann
parents: 33858
diff changeset
  1236
  For \emph{Haskell} a module name prefix may be given using the ``@{text
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1237
  "root:"}'' argument; ``@{text string_classes}'' adds a ``@{verbatim
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1238
  "deriving (Read, Show)"}'' clause to each appropriate datatype
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1239
  declaration.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1240
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1241
  \item @{command (HOL) "code_thms"} prints a list of theorems
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1242
  representing the corresponding program containing all given
34172
4301e9ea1c54 updated documentation
haftmann
parents: 33858
diff changeset
  1243
  constants.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1244
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1245
  \item @{command (HOL) "code_deps"} visualizes dependencies of
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1246
  theorems representing the corresponding program containing all given
34172
4301e9ea1c54 updated documentation
haftmann
parents: 33858
diff changeset
  1247
  constants.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1248
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1249
  \item @{command (HOL) "code_datatype"} specifies a constructor set
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1250
  for a logical type.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1251
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1252
  \item @{command (HOL) "code_const"} associates a list of constants
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1253
  with target-specific serializations; omitting a serialization
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1254
  deletes an existing serialization.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1255
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1256
  \item @{command (HOL) "code_type"} associates a list of type
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1257
  constructors with target-specific serializations; omitting a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1258
  serialization deletes an existing serialization.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1259
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1260
  \item @{command (HOL) "code_class"} associates a list of classes
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1261
  with target-specific class names; omitting a serialization deletes
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1262
  an existing serialization.  This applies only to \emph{Haskell}.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1263
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1264
  \item @{command (HOL) "code_instance"} declares a list of type
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1265
  constructor / class instance relations as ``already present'' for a
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1266
  given target.  Omitting a ``@{text "-"}'' deletes an existing
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1267
  ``already present'' declaration.  This applies only to
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1268
  \emph{Haskell}.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1269
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1270
  \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1271
  to generate monadic code for Haskell.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1272
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1273
  \item @{command (HOL) "code_reserved"} declares a list of names as
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1274
  reserved for a given target, preventing it to be shadowed by any
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1275
  generated code.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1276
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1277
  \item @{command (HOL) "code_include"} adds arbitrary named content
27706
10a6ede68bc8 clarified
haftmann
parents: 27452
diff changeset
  1278
  (``include'') to generated code.  A ``@{text "-"}'' as last argument
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1279
  will remove an already added ``include''.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1280
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1281
  \item @{command (HOL) "code_modulename"} declares aliasings from one
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1282
  module name onto another.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1283
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1284
  \item @{command (HOL) "code_abort"} declares constants which are not
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 29114
diff changeset
  1285
  required to have a definition by means of code equations; if
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1286
  needed these are implemented by program abort instead.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1287
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1288
  \item @{attribute (HOL) code} explicitly selects (or with option
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 29114
diff changeset
  1289
  ``@{text "del"}'' deselects) a code equation for code
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 29114
diff changeset
  1290
  generation.  Usually packages introducing code equations provide
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1291
  a reasonable default setup for selection.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1292
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1293
  \item @{attribute (HOL) code_inline} declares (or with
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27706
diff changeset
  1294
  option ``@{text "del"}'' removes) inlining theorems which are
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 29114
diff changeset
  1295
  applied as rewrite rules to any code equation during
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1296
  preprocessing.
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1297
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1298
  \item @{attribute (HOL) code_post} declares (or with
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1299
  option ``@{text "del"}'' removes) theorems which are
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1300
  applied as rewrite rules to any result of an evaluation.
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31912
diff changeset
  1301
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1302
  \item @{command (HOL) "print_codesetup"} gives an overview on
31254
03a35fbc9dc6 documented print_codeproc command
haftmann
parents: 30865
diff changeset
  1303
  selected code equations and code generator datatypes.
03a35fbc9dc6 documented print_codeproc command
haftmann
parents: 30865
diff changeset
  1304
03a35fbc9dc6 documented print_codeproc command
haftmann
parents: 30865
diff changeset
  1305
  \item @{command (HOL) "print_codeproc"} prints the setup
03a35fbc9dc6 documented print_codeproc command
haftmann
parents: 30865
diff changeset
  1306
  of the code generator preprocessor.
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1307
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1308
  \end{description}
26849
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1309
*}
df50bc1249d7 converted HOL specific elements;
wenzelm
parents: 26840
diff changeset
  1310
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1311
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1312
section {* Definition by specification \label{sec:hol-specification} *}
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1313
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1314
text {*
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1315
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1316
    @{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
  1317
    @{command_def (HOL) "ax_specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1318
  \end{matharray}
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1319
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1320
  \begin{rail}
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1321
  ('specification' | 'ax\_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1322
  ;
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1323
  decl: ((name ':')? term '(' 'overloaded' ')'?)
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1324
  \end{rail}
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1325
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1326
  \begin{description}
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1327
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1328
  \item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1329
  goal stating the existence of terms with the properties specified to
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1330
  hold for the constants given in @{text decls}.  After finishing the
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1331
  proof, the theory will be augmented with definitions for the given
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1332
  constants, as well as with theorems stating the properties for these
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1333
  constants.
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1334
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1335
  \item @{command (HOL) "ax_specification"}~@{text "decls \<phi>"} sets up
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1336
  a goal stating the existence of terms with the properties specified
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1337
  to hold for the constants given in @{text decls}.  After finishing
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1338
  the proof, the theory will be augmented with axioms expressing the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1339
  properties given in the first place.
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1340
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1341
  \item @{text decl} declares a constant to be defined by the
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1342
  specification given.  The definition for the constant @{text c} is
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1343
  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
  1344
  the declaration.  Overloaded constants should be declared as such.
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1345
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28752
diff changeset
  1346
  \end{description}
27045
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1347
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1348
  Whether to use @{command (HOL) "specification"} or @{command (HOL)
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1349
  "ax_specification"} is to some extent a matter of style.  @{command
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1350
  (HOL) "specification"} introduces no new axioms, and so by
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1351
  construction cannot introduce inconsistencies, whereas @{command
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1352
  (HOL) "ax_specification"} does introduce axioms, but only after the
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1353
  user has explicitly proven it to be safe.  A practical issue must be
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1354
  considered, though: After introducing two constants with the same
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1355
  properties using @{command (HOL) "specification"}, one can prove
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1356
  that the two constants are, in fact, equal.  If this might be a
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1357
  problem, one should use @{command (HOL) "ax_specification"}.
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1358
*}
4e7ecec1b685 moved (ax_)specification to end;
wenzelm
parents: 27041
diff changeset
  1359
26840
ec46381f149d added logic-specific sessions;
wenzelm
parents:
diff changeset
  1360
end