| author | wenzelm | 
| Thu, 26 Jul 2012 19:59:06 +0200 | |
| changeset 48536 | 4e2ee88276d2 | 
| parent 48159 | 0b3fd5ff8ea7 | 
| permissions | -rw-r--r-- | 
| 26840 | 1  | 
%  | 
2  | 
\begin{isabellebody}%
 | 
|
| 40406 | 3  | 
\def\isabellecontext{HOL{\isaliteral{5F}{\isacharunderscore}}Specific}%
 | 
| 26840 | 4  | 
%  | 
5  | 
\isadelimtheory  | 
|
6  | 
%  | 
|
7  | 
\endisadelimtheory  | 
|
8  | 
%  | 
|
9  | 
\isatagtheory  | 
|
10  | 
\isacommand{theory}\isamarkupfalse%
 | 
|
| 40406 | 11  | 
\ HOL{\isaliteral{5F}{\isacharunderscore}}Specific\isanewline
 | 
| 44055 | 12  | 
\isakeyword{imports}\ Base\ Main\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7E}{\isachartilde}}{\isaliteral{7E}{\isachartilde}}{\isaliteral{2F}{\isacharslash}}src{\isaliteral{2F}{\isacharslash}}HOL{\isaliteral{2F}{\isacharslash}}Library{\isaliteral{2F}{\isacharslash}}Old{\isaliteral{5F}{\isacharunderscore}}Recdef{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 26849 | 13  | 
\isakeyword{begin}%
 | 
14  | 
\endisatagtheory  | 
|
15  | 
{\isafoldtheory}%
 | 
|
16  | 
%  | 
|
17  | 
\isadelimtheory  | 
|
18  | 
%  | 
|
19  | 
\endisadelimtheory  | 
|
20  | 
%  | 
|
| 26852 | 21  | 
\isamarkupchapter{Isabelle/HOL \label{ch:hol}%
 | 
| 26849 | 22  | 
}  | 
23  | 
\isamarkuptrue%  | 
|
24  | 
%  | 
|
| 42914 | 25  | 
\isamarkupsection{Higher-Order Logic%
 | 
26  | 
}  | 
|
27  | 
\isamarkuptrue%  | 
|
28  | 
%  | 
|
29  | 
\begin{isamarkuptext}%
 | 
|
30  | 
Isabelle/HOL is based on Higher-Order Logic, a polymorphic  | 
|
31  | 
version of Church's Simple Theory of Types. HOL can be best  | 
|
32  | 
understood as a simply-typed version of classical set theory. The  | 
|
33  | 
logic was first implemented in Gordon's HOL system  | 
|
34  | 
  \cite{mgordon-hol}.  It extends Church's original logic
 | 
|
35  | 
  \cite{church40} by explicit type variables (naive polymorphism) and
 | 
|
36  | 
a sound axiomatization scheme for new types based on subsets of  | 
|
37  | 
existing types.  | 
|
38  | 
||
39  | 
  Andrews's book \cite{andrews86} is a full description of the
 | 
|
40  | 
original Church-style higher-order logic, with proofs of correctness  | 
|
41  | 
and completeness wrt.\ certain set-theoretic interpretations. The  | 
|
42  | 
particular extensions of Gordon-style HOL are explained semantically  | 
|
43  | 
  in two chapters of the 1993 HOL book \cite{pitts93}.
 | 
|
44  | 
||
45  | 
Experience with HOL over decades has demonstrated that higher-order  | 
|
46  | 
logic is widely applicable in many areas of mathematics and computer  | 
|
47  | 
science. In a sense, Higher-Order Logic is simpler than First-Order  | 
|
48  | 
Logic, because there are fewer restrictions and special cases. Note  | 
|
49  | 
  that HOL is \emph{weaker} than FOL with axioms for ZF set theory,
 | 
|
50  | 
which is traditionally considered the standard foundation of regular  | 
|
51  | 
mathematics, but for most applications this does not matter. If you  | 
|
52  | 
prefer ML to Lisp, you will probably prefer HOL to ZF.  | 
|
53  | 
||
54  | 
  \medskip The syntax of HOL follows \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-calculus and
 | 
|
55  | 
functional programming. Function application is curried. To apply  | 
|
56  | 
  the function \isa{f} of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{3}}{\isaliteral{22}{\isachardoublequote}}} to the
 | 
|
57  | 
  arguments \isa{a} and \isa{b} in HOL, you simply write \isa{{\isaliteral{22}{\isachardoublequote}}f\ a\ b{\isaliteral{22}{\isachardoublequote}}} (as in ML or Haskell).  There is no ``apply'' operator; the
 | 
|
58  | 
  existing application of the Pure \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-calculus is re-used.
 | 
|
59  | 
  Note that in HOL \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{28}{\isacharparenleft}}a{\isaliteral{2C}{\isacharcomma}}\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} means ``\isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} applied to
 | 
|
60  | 
  the pair \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{2C}{\isacharcomma}}\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} (which is notation for \isa{{\isaliteral{22}{\isachardoublequote}}Pair\ a\ b{\isaliteral{22}{\isachardoublequote}}}).  The latter typically introduces extra formal efforts that can
 | 
|
61  | 
be avoided by currying functions by default. Explicit tuples are as  | 
|
62  | 
infrequent in HOL formalizations as in good ML or Haskell programs.  | 
|
63  | 
||
64  | 
\medskip Isabelle/HOL has a distinct feel, compared to other  | 
|
65  | 
object-logics like Isabelle/ZF. It identifies object-level types  | 
|
66  | 
with meta-level types, taking advantage of the default  | 
|
67  | 
type-inference mechanism of Isabelle/Pure. HOL fully identifies  | 
|
68  | 
object-level functions with meta-level functions, with native  | 
|
69  | 
abstraction and application.  | 
|
70  | 
||
71  | 
These identifications allow Isabelle to support HOL particularly  | 
|
72  | 
nicely, but they also mean that HOL requires some sophistication  | 
|
73  | 
from the user. In particular, an understanding of Hindley-Milner  | 
|
74  | 
type-inference with type-classes, which are both used extensively in  | 
|
75  | 
the standard libraries and applications. Beginners can set  | 
|
76  | 
  \hyperlink{attribute.show-types}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}types}}} or even \hyperlink{attribute.show-sorts}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}sorts}}} to get more
 | 
|
77  | 
explicit information about the result of type-inference.%  | 
|
78  | 
\end{isamarkuptext}%
 | 
|
79  | 
\isamarkuptrue%  | 
|
80  | 
%  | 
|
| 42908 | 81  | 
\isamarkupsection{Inductive and coinductive definitions \label{sec:hol-inductive}%
 | 
| 26849 | 82  | 
}  | 
83  | 
\isamarkuptrue%  | 
|
84  | 
%  | 
|
85  | 
\begin{isamarkuptext}%
 | 
|
| 46280 | 86  | 
\begin{matharray}{rcl}
 | 
87  | 
    \indexdef{HOL}{command}{inductive}\hypertarget{command.HOL.inductive}{\hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
88  | 
    \indexdef{HOL}{command}{inductive\_set}\hypertarget{command.HOL.inductive-set}{\hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
89  | 
    \indexdef{HOL}{command}{coinductive}\hypertarget{command.HOL.coinductive}{\hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
90  | 
    \indexdef{HOL}{command}{coinductive\_set}\hypertarget{command.HOL.coinductive-set}{\hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
91  | 
    \indexdef{HOL}{attribute}{mono}\hypertarget{attribute.HOL.mono}{\hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}}} & : & \isa{attribute} \\
 | 
|
92  | 
  \end{matharray}
 | 
|
93  | 
||
94  | 
  An \emph{inductive definition} specifies the least predicate or set
 | 
|
95  | 
  \isa{R} closed under given rules: applying a rule to elements of
 | 
|
96  | 
  \isa{R} yields a result within \isa{R}.  For example, a
 | 
|
97  | 
structural operational semantics is an inductive definition of an  | 
|
98  | 
evaluation relation.  | 
|
| 42908 | 99  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
100  | 
  Dually, a \emph{coinductive definition} specifies the greatest
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
101  | 
  predicate or set \isa{R} that is consistent with given rules:
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
102  | 
  every element of \isa{R} can be seen as arising by applying a rule
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
103  | 
  to elements of \isa{R}.  An important example is using
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
104  | 
bisimulation relations to formalise equivalence of processes and  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
105  | 
infinite data structures.  | 
| 47859 | 106  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
107  | 
Both inductive and coinductive definitions are based on the  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
108  | 
Knaster-Tarski fixed-point theorem for complete lattices. The  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
109  | 
collection of introduction rules given by the user determines a  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
110  | 
functor on subsets of set-theoretic relations. The required  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
111  | 
monotonicity of the recursion scheme is proven as a prerequisite to  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
112  | 
the fixed-point definition and the resulting consequences. This  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
113  | 
works by pushing inclusion through logical connectives and any other  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
114  | 
operator that might be wrapped around recursive occurrences of the  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
115  | 
defined relation: there must be a monotonicity theorem of the form  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
116  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ B{\isaliteral{22}{\isachardoublequote}}}, for each premise \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4D3E}{\isasymM}}\ R\ t{\isaliteral{22}{\isachardoublequote}}} in an
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
117  | 
introduction rule. The default rule declarations of Isabelle/HOL  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
118  | 
already take care of most common situations.  | 
| 42908 | 119  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
120  | 
  \begin{railoutput}
 | 
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
121  | 
\rail@begin{10}{}
 | 
| 42908 | 122  | 
\rail@bar  | 
123  | 
\rail@term{\hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}}}[]
 | 
|
124  | 
\rail@nextbar{1}
 | 
|
125  | 
\rail@term{\hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}}}[]
 | 
|
126  | 
\rail@nextbar{2}
 | 
|
127  | 
\rail@term{\hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}}[]
 | 
|
128  | 
\rail@nextbar{3}
 | 
|
129  | 
\rail@term{\hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}}}[]
 | 
|
130  | 
\rail@endbar  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
131  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
132  | 
\rail@nextbar{1}
 | 
| 42908 | 133  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
134  | 
\rail@endbar  | 
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
135  | 
\rail@cr{5}
 | 
| 42908 | 136  | 
\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
137  | 
\rail@bar  | 
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
138  | 
\rail@nextbar{6}
 | 
| 42908 | 139  | 
\rail@term{\isa{\isakeyword{for}}}[]
 | 
140  | 
\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[]
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
141  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
142  | 
\rail@bar  | 
| 42908 | 143  | 
\rail@nextbar{6}
 | 
144  | 
\rail@term{\isa{\isakeyword{where}}}[]
 | 
|
145  | 
\rail@nont{\isa{clauses}}[]
 | 
|
146  | 
\rail@endbar  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
147  | 
\rail@cr{8}
 | 
| 42908 | 148  | 
\rail@bar  | 
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
149  | 
\rail@nextbar{9}
 | 
| 42908 | 150  | 
\rail@term{\isa{\isakeyword{monos}}}[]
 | 
151  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
152  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
153  | 
\rail@end  | 
| 42908 | 154  | 
\rail@begin{3}{\isa{clauses}}
 | 
155  | 
\rail@plus  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
156  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
157  | 
\rail@nextbar{1}
 | 
| 42908 | 158  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
159  | 
\rail@endbar  | 
|
160  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
|
161  | 
\rail@nextplus{2}
 | 
|
162  | 
\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[]
 | 
|
163  | 
\rail@endplus  | 
|
164  | 
\rail@end  | 
|
165  | 
\rail@begin{3}{}
 | 
|
166  | 
\rail@term{\hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}}}[]
 | 
|
167  | 
\rail@bar  | 
|
168  | 
\rail@nextbar{1}
 | 
|
169  | 
\rail@term{\isa{add}}[]
 | 
|
170  | 
\rail@nextbar{2}
 | 
|
171  | 
\rail@term{\isa{del}}[]
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
172  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
173  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
174  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
175  | 
|
| 26849 | 176  | 
|
| 28788 | 177  | 
  \begin{description}
 | 
| 42123 | 178  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
179  | 
  \item \hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}} and \hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}} define (co)inductive predicates from the introduction
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
180  | 
rules.  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
181  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
182  | 
  The propositions given as \isa{{\isaliteral{22}{\isachardoublequote}}clauses{\isaliteral{22}{\isachardoublequote}}} in the \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}} part are either rules of the usual \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} format
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
183  | 
  (with arbitrary nesting), or equalities using \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}}.  The
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
184  | 
latter specifies extra-logical abbreviations in the sense of  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
185  | 
  \indexref{}{command}{abbreviation}\hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}.  Introducing abstract syntax
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
186  | 
simultaneously with the actual introduction rules is occasionally  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
187  | 
useful for complex specifications.  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
188  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
189  | 
  The optional \hyperlink{keyword.for}{\mbox{\isa{\isakeyword{for}}}} part contains a list of parameters of
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
190  | 
the (co)inductive predicates that remain fixed throughout the  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
191  | 
definition, in contrast to arguments of the relation that may vary  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
192  | 
  in each occurrence within the given \isa{{\isaliteral{22}{\isachardoublequote}}clauses{\isaliteral{22}{\isachardoublequote}}}.
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
193  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
194  | 
  The optional \hyperlink{keyword.monos}{\mbox{\isa{\isakeyword{monos}}}} declaration contains additional
 | 
| 42908 | 195  | 
  \emph{monotonicity theorems}, which are required for each operator
 | 
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
196  | 
applied to a recursive set in the introduction rules.  | 
| 26849 | 197  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
198  | 
  \item \hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}} and \hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}} are wrappers for to the previous commands for
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
199  | 
native HOL predicates. This allows to define (co)inductive sets,  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
200  | 
where multiple arguments are simulated via tuples.  | 
| 26849 | 201  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
202  | 
  \item \hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}} declares monotonicity rules in the
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
203  | 
context. These rule are involved in the automated monotonicity  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
204  | 
proof of the above inductive and coinductive definitions.  | 
| 26849 | 205  | 
|
| 28788 | 206  | 
  \end{description}%
 | 
| 26849 | 207  | 
\end{isamarkuptext}%
 | 
208  | 
\isamarkuptrue%  | 
|
209  | 
%  | 
|
| 42908 | 210  | 
\isamarkupsubsection{Derived rules%
 | 
| 26849 | 211  | 
}  | 
212  | 
\isamarkuptrue%  | 
|
213  | 
%  | 
|
214  | 
\begin{isamarkuptext}%
 | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
215  | 
A (co)inductive definition of \isa{R} provides the following
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
216  | 
main theorems:  | 
| 26849 | 217  | 
|
| 42908 | 218  | 
  \begin{description}
 | 
| 26849 | 219  | 
|
| 42908 | 220  | 
  \item \isa{R{\isaliteral{2E}{\isachardot}}intros} is the list of introduction rules as proven
 | 
221  | 
theorems, for the recursive predicates (or sets). The rules are  | 
|
222  | 
also available individually, using the names given them in the  | 
|
223  | 
theory file;  | 
|
| 26849 | 224  | 
|
| 42908 | 225  | 
  \item \isa{R{\isaliteral{2E}{\isachardot}}cases} is the case analysis (or elimination) rule;
 | 
226  | 
||
227  | 
  \item \isa{R{\isaliteral{2E}{\isachardot}}induct} or \isa{R{\isaliteral{2E}{\isachardot}}coinduct} is the (co)induction
 | 
|
| 
44273
 
336752fb25df
adding documentation about simps equation in the inductive package
 
bulwahn 
parents: 
44055 
diff
changeset
 | 
228  | 
rule;  | 
| 
 
336752fb25df
adding documentation about simps equation in the inductive package
 
bulwahn 
parents: 
44055 
diff
changeset
 | 
229  | 
|
| 
 
336752fb25df
adding documentation about simps equation in the inductive package
 
bulwahn 
parents: 
44055 
diff
changeset
 | 
230  | 
  \item \isa{R{\isaliteral{2E}{\isachardot}}simps} is the equation unrolling the fixpoint of the
 | 
| 
 
336752fb25df
adding documentation about simps equation in the inductive package
 
bulwahn 
parents: 
44055 
diff
changeset
 | 
231  | 
predicate one step.  | 
| 47859 | 232  | 
|
| 42908 | 233  | 
  \end{description}
 | 
| 26849 | 234  | 
|
| 42908 | 235  | 
  When several predicates \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} are
 | 
236  | 
defined simultaneously, the list of introduction rules is called  | 
|
237  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5F}{\isacharunderscore}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}intros{\isaliteral{22}{\isachardoublequote}}}, the case analysis rules are
 | 
|
238  | 
  called \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2E}{\isachardot}}cases{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}cases{\isaliteral{22}{\isachardoublequote}}}, and the list
 | 
|
239  | 
  of mutual induction rules is called \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5F}{\isacharunderscore}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}inducts{\isaliteral{22}{\isachardoublequote}}}.%
 | 
|
| 26849 | 240  | 
\end{isamarkuptext}%
 | 
241  | 
\isamarkuptrue%  | 
|
242  | 
%  | 
|
| 42908 | 243  | 
\isamarkupsubsection{Monotonicity theorems%
 | 
| 26849 | 244  | 
}  | 
245  | 
\isamarkuptrue%  | 
|
246  | 
%  | 
|
247  | 
\begin{isamarkuptext}%
 | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
248  | 
The context maintains a default set of theorems that are used  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
249  | 
in monotonicity proofs. New rules can be declared via the  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
250  | 
  \hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}} attribute.  See the main Isabelle/HOL
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
251  | 
sources for some examples. The general format of such monotonicity  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
252  | 
theorems is as follows:  | 
| 26849 | 253  | 
|
| 42908 | 254  | 
  \begin{itemize}
 | 
| 26849 | 255  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
256  | 
  \item Theorems of the form \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ B{\isaliteral{22}{\isachardoublequote}}}, for proving
 | 
| 42908 | 257  | 
monotonicity of inductive definitions whose introduction rules have  | 
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
258  | 
  premises involving terms such as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4D3E}{\isasymM}}\ R\ t{\isaliteral{22}{\isachardoublequote}}}.
 | 
| 26849 | 259  | 
|
| 42908 | 260  | 
\item Monotonicity theorems for logical operators, which are of the  | 
261  | 
  general form \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}.  For example, in
 | 
|
262  | 
  the case of the operator \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6F723E}{\isasymor}}{\isaliteral{22}{\isachardoublequote}}}, the corresponding theorem is
 | 
|
263  | 
\[  | 
|
264  | 
  \infer{\isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}}
 | 
|
265  | 
\]  | 
|
| 41396 | 266  | 
|
| 42908 | 267  | 
\item De Morgan style equations for reasoning about the ``polarity''  | 
268  | 
of expressions, e.g.  | 
|
269  | 
\[  | 
|
270  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ P{\isaliteral{22}{\isachardoublequote}}} \qquad\qquad
 | 
|
271  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q{\isaliteral{22}{\isachardoublequote}}}
 | 
|
272  | 
\]  | 
|
| 41396 | 273  | 
|
| 42908 | 274  | 
\item Equations for reducing complex operators to more primitive  | 
275  | 
ones whose monotonicity can easily be proved, e.g.  | 
|
276  | 
\[  | 
|
277  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}} \qquad\qquad
 | 
|
278  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}Ball\ A\ P\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ x{\isaliteral{22}{\isachardoublequote}}}
 | 
|
279  | 
\]  | 
|
| 41396 | 280  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
281  | 
  \end{itemize}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
282  | 
\end{isamarkuptext}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
283  | 
\isamarkuptrue%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
284  | 
%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
285  | 
\isamarkupsubsubsection{Examples%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
286  | 
}  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
287  | 
\isamarkuptrue%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
288  | 
%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
289  | 
\begin{isamarkuptext}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
290  | 
The finite powerset operator can be defined inductively like this:%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
291  | 
\end{isamarkuptext}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
292  | 
\isamarkuptrue%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
293  | 
\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}\isamarkupfalse%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
294  | 
\ Fin\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ set\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ set\ set{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{for}\ A\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
295  | 
\isakeyword{where}\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
296  | 
\ \ empty{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ Fin\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
297  | 
{\isaliteral{7C}{\isacharbar}}\ insert{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C696E3E}{\isasymin}}\ Fin\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ insert\ a\ B\ {\isaliteral{5C3C696E3E}{\isasymin}}\ Fin\ A{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
298  | 
\begin{isamarkuptext}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
299  | 
The accessible part of a relation is defined as follows:%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
300  | 
\end{isamarkuptext}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
301  | 
\isamarkuptrue%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
302  | 
\isacommand{inductive}\isamarkupfalse%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
303  | 
\ acc\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
304  | 
\ \ \isakeyword{for}\ r\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infix}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C707265633E}{\isasymprec}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
305  | 
\isakeyword{where}\ acc{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}y{\isaliteral{2E}{\isachardot}}\ y\ {\isaliteral{5C3C707265633E}{\isasymprec}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ acc\ r\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ acc\ r\ x{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
306  | 
\begin{isamarkuptext}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
307  | 
Common logical connectives can be easily characterized as  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
308  | 
non-recursive inductive definitions with parameters, but without  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
309  | 
arguments.%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
310  | 
\end{isamarkuptext}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
311  | 
\isamarkuptrue%  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
312  | 
\isacommand{inductive}\isamarkupfalse%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
313  | 
\ AND\ \isakeyword{for}\ A\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ bool\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
314  | 
\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ AND\ A\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
315  | 
\isanewline  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
316  | 
\isacommand{inductive}\isamarkupfalse%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
317  | 
\ OR\ \isakeyword{for}\ A\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ bool\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
318  | 
\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ OR\ A\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
319  | 
\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ OR\ A\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
320  | 
\isanewline  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
321  | 
\isacommand{inductive}\isamarkupfalse%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
322  | 
\ EXISTS\ \isakeyword{for}\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
323  | 
\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ a\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ EXISTS\ B{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
324  | 
\begin{isamarkuptext}%
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
325  | 
Here the \isa{{\isaliteral{22}{\isachardoublequote}}cases{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}induct{\isaliteral{22}{\isachardoublequote}}} rules produced by
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
326  | 
  the \hyperlink{command.inductive}{\mbox{\isa{\isacommand{inductive}}}} package coincide with the expected
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
327  | 
elimination rules for Natural Deduction. Already in the original  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
328  | 
  article by Gerhard Gentzen \cite{Gentzen:1935} there is a hint that
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
329  | 
each connective can be characterized by its introductions, and the  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
330  | 
elimination can be constructed systematically.%  | 
| 41396 | 331  | 
\end{isamarkuptext}%
 | 
332  | 
\isamarkuptrue%  | 
|
333  | 
%  | 
|
| 26849 | 334  | 
\isamarkupsection{Recursive functions \label{sec:recursion}%
 | 
335  | 
}  | 
|
336  | 
\isamarkuptrue%  | 
|
337  | 
%  | 
|
338  | 
\begin{isamarkuptext}%
 | 
|
339  | 
\begin{matharray}{rcl}
 | 
|
| 40406 | 340  | 
    \indexdef{HOL}{command}{primrec}\hypertarget{command.HOL.primrec}{\hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
341  | 
    \indexdef{HOL}{command}{fun}\hypertarget{command.HOL.fun}{\hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
342  | 
    \indexdef{HOL}{command}{function}\hypertarget{command.HOL.function}{\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
343  | 
    \indexdef{HOL}{command}{termination}\hypertarget{command.HOL.termination}{\hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 26849 | 344  | 
  \end{matharray}
 | 
345  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
346  | 
  \begin{railoutput}
 | 
| 42662 | 347  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
348  | 
\rail@term{\hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
349  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
350  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
351  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
352  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
353  | 
\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
354  | 
\rail@term{\isa{\isakeyword{where}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
355  | 
\rail@nont{\isa{equations}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
356  | 
\rail@end  | 
| 42662 | 357  | 
\rail@begin{4}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
358  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
359  | 
\rail@term{\hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
360  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
361  | 
\rail@term{\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
362  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
363  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
364  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
365  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
366  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
367  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
368  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
369  | 
\rail@nont{\isa{functionopts}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
370  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
371  | 
\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
372  | 
\rail@cr{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
373  | 
\rail@term{\isa{\isakeyword{where}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
374  | 
\rail@nont{\isa{equations}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
375  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
376  | 
\rail@begin{3}{\isa{equations}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
377  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
378  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
379  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
380  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
381  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
382  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
383  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
384  | 
\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
385  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
386  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
387  | 
\rail@begin{3}{\isa{functionopts}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
388  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
389  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
390  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
391  | 
\rail@term{\isa{sequential}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
392  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
393  | 
\rail@term{\isa{domintros}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
394  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
395  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
396  | 
\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
397  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
398  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
399  | 
\rail@end  | 
| 42662 | 400  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
401  | 
\rail@term{\hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
402  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
403  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
404  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
405  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
406  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
407  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
408  | 
|
| 26849 | 409  | 
|
| 28788 | 410  | 
  \begin{description}
 | 
| 26849 | 411  | 
|
| 28788 | 412  | 
  \item \hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}} defines primitive recursive
 | 
| 42912 | 413  | 
  functions over datatypes (see also \indexref{HOL}{command}{datatype}\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}} and
 | 
414  | 
  \indexref{HOL}{command}{rep\_datatype}\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}).  The given \isa{equations}
 | 
|
415  | 
specify reduction rules that are produced by instantiating the  | 
|
416  | 
generic combinator for primitive recursion that is available for  | 
|
417  | 
each datatype.  | 
|
418  | 
||
419  | 
Each equation needs to be of the form:  | 
|
420  | 
||
421  | 
  \begin{isabelle}%
 | 
|
422  | 
{\isaliteral{22}{\isachardoublequote}}f\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{28}{\isacharparenleft}}C\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{29}{\isacharparenright}}\ z\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ z\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3D}{\isacharequal}}\ rhs{\isaliteral{22}{\isachardoublequote}}%
 | 
|
423  | 
\end{isabelle}
 | 
|
424  | 
||
425  | 
  such that \isa{C} is a datatype constructor, \isa{rhs} contains
 | 
|
426  | 
only the free variables on the left-hand side (or from the context),  | 
|
427  | 
  and all recursive occurrences of \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} in \isa{{\isaliteral{22}{\isachardoublequote}}rhs{\isaliteral{22}{\isachardoublequote}}} are of
 | 
|
428  | 
  the form \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} for some \isa{i}.  At most one
 | 
|
429  | 
reduction rule for each constructor can be given. The order does  | 
|
430  | 
not matter. For missing constructors, the function is defined to  | 
|
431  | 
return a default value, but this equation is made difficult to  | 
|
432  | 
access for users.  | 
|
433  | 
||
434  | 
  The reduction rules are declared as \hyperlink{attribute.simp}{\mbox{\isa{simp}}} by default,
 | 
|
435  | 
  which enables standard proof methods like \hyperlink{method.simp}{\mbox{\isa{simp}}} and
 | 
|
436  | 
  \hyperlink{method.auto}{\mbox{\isa{auto}}} to normalize expressions of \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} applied to
 | 
|
437  | 
datatype constructions, by simulating symbolic computation via  | 
|
438  | 
rewriting.  | 
|
| 26849 | 439  | 
|
| 28788 | 440  | 
  \item \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} defines functions by general
 | 
| 26849 | 441  | 
wellfounded recursion. A detailed description with examples can be  | 
442  | 
  found in \cite{isabelle-function}. The function is specified by a
 | 
|
443  | 
set of (possibly conditional) recursive equations with arbitrary  | 
|
444  | 
pattern matching. The command generates proof obligations for the  | 
|
445  | 
completeness and the compatibility of patterns.  | 
|
446  | 
||
447  | 
The defined function is considered partial, and the resulting  | 
|
| 40406 | 448  | 
  simplification rules (named \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}psimps{\isaliteral{22}{\isachardoublequote}}}) and induction rule
 | 
449  | 
  (named \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}pinduct{\isaliteral{22}{\isachardoublequote}}}) are guarded by a generated domain
 | 
|
450  | 
  predicate \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{5F}{\isacharunderscore}}dom{\isaliteral{22}{\isachardoublequote}}}. The \hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}
 | 
|
| 26849 | 451  | 
command can then be used to establish that the function is total.  | 
452  | 
||
| 40406 | 453  | 
  \item \hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}} is a shorthand notation for ``\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}sequential{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, followed by automated
 | 
| 28788 | 454  | 
proof attempts regarding pattern matching and termination. See  | 
455  | 
  \cite{isabelle-function} for further details.
 | 
|
| 26849 | 456  | 
|
| 28788 | 457  | 
  \item \hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}~\isa{f} commences a
 | 
| 26849 | 458  | 
  termination proof for the previously defined function \isa{f}.  If
 | 
459  | 
this is omitted, the command refers to the most recent function  | 
|
460  | 
definition. After the proof is closed, the recursive equations and  | 
|
461  | 
the induction principle is established.  | 
|
462  | 
||
| 28788 | 463  | 
  \end{description}
 | 
| 26849 | 464  | 
|
| 27452 | 465  | 
  Recursive definitions introduced by the \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}
 | 
| 42912 | 466  | 
  command accommodate reasoning by induction (cf.\ \hyperlink{method.induct}{\mbox{\isa{induct}}}):
 | 
467  | 
  rule \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}induct{\isaliteral{22}{\isachardoublequote}}} refers to a specific induction rule, with
 | 
|
468  | 
parameters named according to the user-specified equations. Cases  | 
|
469  | 
  are numbered starting from 1.  For \hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}, the
 | 
|
470  | 
induction principle coincides with structural recursion on the  | 
|
471  | 
datatype where the recursion is carried out.  | 
|
| 26849 | 472  | 
|
473  | 
The equations provided by these packages may be referred later as  | 
|
| 40406 | 474  | 
  theorem list \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}simps{\isaliteral{22}{\isachardoublequote}}}, where \isa{f} is the (collective)
 | 
| 26849 | 475  | 
name of the functions defined. Individual equations may be named  | 
476  | 
explicitly as well.  | 
|
477  | 
||
| 26902 | 478  | 
  The \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} command accepts the following
 | 
| 26849 | 479  | 
options.  | 
480  | 
||
| 28788 | 481  | 
  \begin{description}
 | 
| 26849 | 482  | 
|
| 28788 | 483  | 
  \item \isa{sequential} enables a preprocessor which disambiguates
 | 
484  | 
overlapping patterns by making them mutually disjoint. Earlier  | 
|
485  | 
equations take precedence over later ones. This allows to give the  | 
|
486  | 
specification in a format very similar to functional programming.  | 
|
487  | 
Note that the resulting simplification and induction rules  | 
|
488  | 
correspond to the transformed specification, not the one given  | 
|
| 26849 | 489  | 
originally. This usually means that each equation given by the user  | 
| 36139 | 490  | 
may result in several theorems. Also note that this automatic  | 
| 26849 | 491  | 
transformation only works for ML-style datatype patterns.  | 
492  | 
||
| 28788 | 493  | 
  \item \isa{domintros} enables the automated generation of
 | 
| 26849 | 494  | 
introduction rules for the domain predicate. While mostly not  | 
495  | 
needed, they can be helpful in some proofs about partial functions.  | 
|
496  | 
||
| 28788 | 497  | 
  \end{description}%
 | 
| 26849 | 498  | 
\end{isamarkuptext}%
 | 
499  | 
\isamarkuptrue%  | 
|
500  | 
%  | 
|
| 42912 | 501  | 
\isamarkupsubsubsection{Example: evaluation of expressions%
 | 
502  | 
}  | 
|
503  | 
\isamarkuptrue%  | 
|
504  | 
%  | 
|
505  | 
\begin{isamarkuptext}%
 | 
|
506  | 
Subsequently, we define mutual datatypes for arithmetic and  | 
|
507  | 
  boolean expressions, and use \hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}} for evaluation
 | 
|
508  | 
functions that follow the same recursive structure.%  | 
|
509  | 
\end{isamarkuptext}%
 | 
|
510  | 
\isamarkuptrue%  | 
|
511  | 
\isacommand{datatype}\isamarkupfalse%
 | 
|
512  | 
\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{3D}{\isacharequal}}\isanewline
 | 
|
513  | 
\ \ \ \ IF\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
514  | 
\ \ {\isaliteral{7C}{\isacharbar}}\ Sum\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
515  | 
\ \ {\isaliteral{7C}{\isacharbar}}\ Diff\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
516  | 
\ \ {\isaliteral{7C}{\isacharbar}}\ Var\ {\isaliteral{27}{\isacharprime}}a\isanewline
 | 
|
517  | 
\ \ {\isaliteral{7C}{\isacharbar}}\ Num\ nat\isanewline
 | 
|
518  | 
\isakeyword{and}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{3D}{\isacharequal}}\isanewline
 | 
|
519  | 
\ \ \ \ Less\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
520  | 
\ \ {\isaliteral{7C}{\isacharbar}}\ And\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
521  | 
\ \ {\isaliteral{7C}{\isacharbar}}\ Neg\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
|
522  | 
\begin{isamarkuptext}%
 | 
|
523  | 
\medskip Evaluation of arithmetic and boolean expressions%  | 
|
524  | 
\end{isamarkuptext}%
 | 
|
525  | 
\isamarkuptrue%  | 
|
526  | 
\isacommand{primrec}\isamarkupfalse%
 | 
|
527  | 
\ evala\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
528  | 
\ \ \isakeyword{and}\ evalb\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
529  | 
\isakeyword{where}\isanewline
 | 
|
530  | 
\ \ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ evalb\ env\ b\ then\ evala\ env\ a{\isadigit{1}}\ else\ evala\ env\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
531  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Sum\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isadigit{1}}\ {\isaliteral{2B}{\isacharplus}}\ evala\ env\ a{\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
532  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Diff\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isadigit{1}}\ {\isaliteral{2D}{\isacharminus}}\ evala\ env\ a{\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
533  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Var\ v{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ env\ v{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
534  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Num\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
535  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}evala\ env\ a{\isadigit{1}}\ {\isaliteral{3C}{\isacharless}}\ evala\ env\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
536  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}evalb\ env\ b{\isadigit{1}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ evalb\ env\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
537  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}Neg\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ evalb\ env\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
|
538  | 
\begin{isamarkuptext}%
 | 
|
539  | 
Since the value of an expression depends on the value of its  | 
|
540  | 
  variables, the functions \isa{evala} and \isa{evalb} take an
 | 
|
541  | 
  additional parameter, an \emph{environment} that maps variables to
 | 
|
542  | 
their values.  | 
|
543  | 
||
544  | 
\medskip Substitution on expressions can be defined similarly. The  | 
|
545  | 
  mapping \isa{f} of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequote}}} given as a
 | 
|
546  | 
  parameter is lifted canonically on the types \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequote}}} and
 | 
|
547  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequote}}}, respectively.%
 | 
|
548  | 
\end{isamarkuptext}%
 | 
|
549  | 
\isamarkuptrue%  | 
|
550  | 
\isacommand{primrec}\isamarkupfalse%
 | 
|
551  | 
\ substa\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
552  | 
\ \ \isakeyword{and}\ substb\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
553  | 
\isakeyword{where}\isanewline
 | 
|
554  | 
\ \ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ IF\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
555  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Sum\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Sum\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
556  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Diff\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Diff\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
557  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Var\ v{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ f\ v{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
558  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Num\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Num\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
559  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substb\ f\ {\isaliteral{28}{\isacharparenleft}}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Less\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
560  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substb\ f\ {\isaliteral{28}{\isacharparenleft}}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ And\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
561  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substb\ f\ {\isaliteral{28}{\isacharparenleft}}Neg\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Neg\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
|
562  | 
\begin{isamarkuptext}%
 | 
|
563  | 
In textbooks about semantics one often finds substitution  | 
|
564  | 
theorems, which express the relationship between substitution and  | 
|
565  | 
  evaluation.  For \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequote}}}, we can prove
 | 
|
566  | 
such a theorem by mutual induction, followed by simplification.%  | 
|
567  | 
\end{isamarkuptext}%
 | 
|
568  | 
\isamarkuptrue%  | 
|
569  | 
\isacommand{lemma}\isamarkupfalse%
 | 
|
570  | 
\ subst{\isaliteral{5F}{\isacharunderscore}}one{\isaliteral{3A}{\isacharcolon}}\isanewline
 | 
|
571  | 
\ \ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}substa\ {\isaliteral{28}{\isacharparenleft}}Var\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ {\isaliteral{28}{\isacharparenleft}}env\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
572  | 
\ \ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}substb\ {\isaliteral{28}{\isacharparenleft}}Var\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evalb\ {\isaliteral{28}{\isacharparenleft}}env\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
573  | 
%  | 
|
574  | 
\isadelimproof  | 
|
575  | 
\ \ %  | 
|
576  | 
\endisadelimproof  | 
|
577  | 
%  | 
|
578  | 
\isatagproof  | 
|
579  | 
\isacommand{by}\isamarkupfalse%
 | 
|
580  | 
\ {\isaliteral{28}{\isacharparenleft}}induct\ a\ \isakeyword{and}\ b{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all%
 | 
|
581  | 
\endisatagproof  | 
|
582  | 
{\isafoldproof}%
 | 
|
583  | 
%  | 
|
584  | 
\isadelimproof  | 
|
585  | 
\isanewline  | 
|
586  | 
%  | 
|
587  | 
\endisadelimproof  | 
|
588  | 
\isanewline  | 
|
589  | 
\isacommand{lemma}\isamarkupfalse%
 | 
|
590  | 
\ subst{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{3A}{\isacharcolon}}\isanewline
 | 
|
591  | 
\ \ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ evala\ env\ {\isaliteral{28}{\isacharparenleft}}s\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
592  | 
\ \ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}substb\ s\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evalb\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ evala\ env\ {\isaliteral{28}{\isacharparenleft}}s\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
593  | 
%  | 
|
594  | 
\isadelimproof  | 
|
595  | 
\ \ %  | 
|
596  | 
\endisadelimproof  | 
|
597  | 
%  | 
|
598  | 
\isatagproof  | 
|
599  | 
\isacommand{by}\isamarkupfalse%
 | 
|
600  | 
\ {\isaliteral{28}{\isacharparenleft}}induct\ a\ \isakeyword{and}\ b{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all%
 | 
|
601  | 
\endisatagproof  | 
|
602  | 
{\isafoldproof}%
 | 
|
603  | 
%  | 
|
604  | 
\isadelimproof  | 
|
605  | 
%  | 
|
606  | 
\endisadelimproof  | 
|
607  | 
%  | 
|
608  | 
\isamarkupsubsubsection{Example: a substitution function for terms%
 | 
|
609  | 
}  | 
|
610  | 
\isamarkuptrue%  | 
|
611  | 
%  | 
|
612  | 
\begin{isamarkuptext}%
 | 
|
613  | 
Functions on datatypes with nested recursion are also defined  | 
|
614  | 
by mutual primitive recursion.%  | 
|
615  | 
\end{isamarkuptext}%
 | 
|
616  | 
\isamarkuptrue%  | 
|
617  | 
\isacommand{datatype}\isamarkupfalse%
 | 
|
618  | 
\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{22}{\isachardoublequoteopen}}term{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{3D}{\isacharequal}}\ Var\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{7C}{\isacharbar}}\ App\ {\isaliteral{27}{\isacharprime}}b\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
|
619  | 
\begin{isamarkuptext}%
 | 
|
620  | 
A substitution function on type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{22}{\isachardoublequote}}} can be
 | 
|
621  | 
  defined as follows, by working simultaneously on \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list{\isaliteral{22}{\isachardoublequote}}}:%
 | 
|
622  | 
\end{isamarkuptext}%
 | 
|
623  | 
\isamarkuptrue%  | 
|
624  | 
\isacommand{primrec}\isamarkupfalse%
 | 
|
625  | 
\ subst{\isaliteral{5F}{\isacharunderscore}}term\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline
 | 
|
626  | 
\ \ subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
627  | 
\isakeyword{where}\isanewline
 | 
|
628  | 
\ \ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f\ {\isaliteral{28}{\isacharparenleft}}Var\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ f\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
629  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f\ {\isaliteral{28}{\isacharparenleft}}App\ b\ ts{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ App\ b\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ ts{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
630  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
631  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ {\isaliteral{28}{\isacharparenleft}}t\ {\isaliteral{23}{\isacharhash}}\ ts{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ subst{\isaliteral{5F}{\isacharunderscore}}term\ f\ t\ {\isaliteral{23}{\isacharhash}}\ subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ ts{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
|
632  | 
\begin{isamarkuptext}%
 | 
|
633  | 
The recursion scheme follows the structure of the unfolded  | 
|
634  | 
  definition of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{22}{\isachardoublequote}}}.  To prove properties of this
 | 
|
635  | 
substitution function, mutual induction is needed:%  | 
|
636  | 
\end{isamarkuptext}%
 | 
|
637  | 
\isamarkuptrue%  | 
|
638  | 
\isacommand{lemma}\isamarkupfalse%
 | 
|
639  | 
\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ f{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{1}}\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{2}}\ t{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline
 | 
|
640  | 
\ \ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ f{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ ts\ {\isaliteral{3D}{\isacharequal}}\ subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f{\isadigit{1}}\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f{\isadigit{2}}\ ts{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
641  | 
%  | 
|
642  | 
\isadelimproof  | 
|
643  | 
\ \ %  | 
|
644  | 
\endisadelimproof  | 
|
645  | 
%  | 
|
646  | 
\isatagproof  | 
|
647  | 
\isacommand{by}\isamarkupfalse%
 | 
|
648  | 
\ {\isaliteral{28}{\isacharparenleft}}induct\ t\ \isakeyword{and}\ ts{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all%
 | 
|
649  | 
\endisatagproof  | 
|
650  | 
{\isafoldproof}%
 | 
|
651  | 
%  | 
|
652  | 
\isadelimproof  | 
|
653  | 
%  | 
|
654  | 
\endisadelimproof  | 
|
655  | 
%  | 
|
656  | 
\isamarkupsubsubsection{Example: a map function for infinitely branching trees%
 | 
|
657  | 
}  | 
|
658  | 
\isamarkuptrue%  | 
|
659  | 
%  | 
|
660  | 
\begin{isamarkuptext}%
 | 
|
661  | 
Defining functions on infinitely branching datatypes by  | 
|
662  | 
primitive recursion is just as easy.%  | 
|
663  | 
\end{isamarkuptext}%
 | 
|
664  | 
\isamarkuptrue%  | 
|
665  | 
\isacommand{datatype}\isamarkupfalse%
 | 
|
666  | 
\ {\isaliteral{27}{\isacharprime}}a\ tree\ {\isaliteral{3D}{\isacharequal}}\ Atom\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{7C}{\isacharbar}}\ Branch\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ tree{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
667  | 
\isanewline  | 
|
668  | 
\isacommand{primrec}\isamarkupfalse%
 | 
|
669  | 
\ map{\isaliteral{5F}{\isacharunderscore}}tree\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ tree\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ tree{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
670  | 
\isakeyword{where}\isanewline
 | 
|
671  | 
\ \ {\isaliteral{22}{\isachardoublequoteopen}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{28}{\isacharparenleft}}Atom\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Atom\ {\isaliteral{28}{\isacharparenleft}}f\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
672  | 
{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{28}{\isacharparenleft}}Branch\ ts{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Branch\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{28}{\isacharparenleft}}ts\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
|
673  | 
\begin{isamarkuptext}%
 | 
|
674  | 
Note that all occurrences of functions such as \isa{ts}
 | 
|
675  | 
  above must be applied to an argument.  In particular, \isa{{\isaliteral{22}{\isachardoublequote}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ ts{\isaliteral{22}{\isachardoublequote}}} is not allowed here.%
 | 
|
676  | 
\end{isamarkuptext}%
 | 
|
677  | 
\isamarkuptrue%  | 
|
678  | 
%  | 
|
679  | 
\begin{isamarkuptext}%
 | 
|
680  | 
Here is a simple composition lemma for \isa{map{\isaliteral{5F}{\isacharunderscore}}tree}:%
 | 
|
681  | 
\end{isamarkuptext}%
 | 
|
682  | 
\isamarkuptrue%  | 
|
683  | 
\isacommand{lemma}\isamarkupfalse%
 | 
|
684  | 
\ {\isaliteral{22}{\isachardoublequoteopen}}map{\isaliteral{5F}{\isacharunderscore}}tree\ g\ {\isaliteral{28}{\isacharparenleft}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ t{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ map{\isaliteral{5F}{\isacharunderscore}}tree\ {\isaliteral{28}{\isacharparenleft}}g\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ f{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
685  | 
%  | 
|
686  | 
\isadelimproof  | 
|
687  | 
\ \ %  | 
|
688  | 
\endisadelimproof  | 
|
689  | 
%  | 
|
690  | 
\isatagproof  | 
|
691  | 
\isacommand{by}\isamarkupfalse%
 | 
|
692  | 
\ {\isaliteral{28}{\isacharparenleft}}induct\ t{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all%
 | 
|
693  | 
\endisatagproof  | 
|
694  | 
{\isafoldproof}%
 | 
|
695  | 
%  | 
|
696  | 
\isadelimproof  | 
|
697  | 
%  | 
|
698  | 
\endisadelimproof  | 
|
699  | 
%  | 
|
| 26849 | 700  | 
\isamarkupsubsection{Proof methods related to recursive definitions%
 | 
701  | 
}  | 
|
702  | 
\isamarkuptrue%  | 
|
703  | 
%  | 
|
704  | 
\begin{isamarkuptext}%
 | 
|
705  | 
\begin{matharray}{rcl}
 | 
|
| 40406 | 706  | 
    \indexdef{HOL}{method}{pat\_completeness}\hypertarget{method.HOL.pat-completeness}{\hyperlink{method.HOL.pat-completeness}{\mbox{\isa{pat{\isaliteral{5F}{\isacharunderscore}}completeness}}}} & : & \isa{method} \\
 | 
| 28788 | 707  | 
    \indexdef{HOL}{method}{relation}\hypertarget{method.HOL.relation}{\hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}} & : & \isa{method} \\
 | 
| 40406 | 708  | 
    \indexdef{HOL}{method}{lexicographic\_order}\hypertarget{method.HOL.lexicographic-order}{\hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}}} & : & \isa{method} \\
 | 
709  | 
    \indexdef{HOL}{method}{size\_change}\hypertarget{method.HOL.size-change}{\hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}}} & : & \isa{method} \\
 | 
|
| 
45944
 
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
 
bulwahn 
parents: 
45943 
diff
changeset
 | 
710  | 
    \indexdef{HOL}{method}{induction\_schema}\hypertarget{method.HOL.induction-schema}{\hyperlink{method.HOL.induction-schema}{\mbox{\isa{induction{\isaliteral{5F}{\isacharunderscore}}schema}}}} & : & \isa{method} \\
 | 
| 26849 | 711  | 
  \end{matharray}
 | 
712  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
713  | 
  \begin{railoutput}
 | 
| 42662 | 714  | 
\rail@begin{1}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
715  | 
\rail@term{\hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
716  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
717  | 
\rail@end  | 
| 42662 | 718  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
719  | 
\rail@term{\hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
720  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
721  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
722  | 
\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
723  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
724  | 
\rail@end  | 
| 42662 | 725  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
726  | 
\rail@term{\hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
727  | 
\rail@nont{\isa{orders}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
728  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
729  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
730  | 
\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
731  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
732  | 
\rail@end  | 
| 
45944
 
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
 
bulwahn 
parents: 
45943 
diff
changeset
 | 
733  | 
\rail@begin{1}{}
 | 
| 
 
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
 
bulwahn 
parents: 
45943 
diff
changeset
 | 
734  | 
\rail@term{\hyperlink{method.HOL.induction-schema}{\mbox{\isa{induction{\isaliteral{5F}{\isacharunderscore}}schema}}}}[]
 | 
| 
 
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
 
bulwahn 
parents: 
45943 
diff
changeset
 | 
735  | 
\rail@end  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
736  | 
\rail@begin{4}{\isa{orders}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
737  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
738  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
739  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
740  | 
\rail@term{\isa{max}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
741  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
742  | 
\rail@term{\isa{min}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
743  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
744  | 
\rail@term{\isa{ms}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
745  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
746  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
747  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
748  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
749  | 
|
| 26849 | 750  | 
|
| 28788 | 751  | 
  \begin{description}
 | 
| 26849 | 752  | 
|
| 40406 | 753  | 
  \item \hyperlink{method.HOL.pat-completeness}{\mbox{\isa{pat{\isaliteral{5F}{\isacharunderscore}}completeness}}} is a specialized method to
 | 
| 26849 | 754  | 
solve goals regarding the completeness of pattern matching, as  | 
| 26902 | 755  | 
  required by the \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} package (cf.\
 | 
| 26849 | 756  | 
  \cite{isabelle-function}).
 | 
757  | 
||
| 28788 | 758  | 
  \item \hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}~\isa{R} introduces a termination
 | 
| 26849 | 759  | 
  proof using the relation \isa{R}.  The resulting proof state will
 | 
760  | 
  contain goals expressing that \isa{R} is wellfounded, and that the
 | 
|
761  | 
  arguments of recursive calls decrease with respect to \isa{R}.
 | 
|
762  | 
Usually, this method is used as the initial proof step of manual  | 
|
763  | 
termination proofs.  | 
|
764  | 
||
| 40406 | 765  | 
  \item \hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}} attempts a fully
 | 
| 26849 | 766  | 
automated termination proof by searching for a lexicographic  | 
767  | 
combination of size measures on the arguments of the function. The  | 
|
| 26902 | 768  | 
  method accepts the same arguments as the \hyperlink{method.auto}{\mbox{\isa{auto}}} method,
 | 
| 42930 | 769  | 
  which it uses internally to prove local descents.  The \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}} modifiers are accepted (as for \hyperlink{method.auto}{\mbox{\isa{auto}}}).
 | 
| 26849 | 770  | 
|
771  | 
In case of failure, extensive information is printed, which can help  | 
|
772  | 
  to analyse the situation (cf.\ \cite{isabelle-function}).
 | 
|
773  | 
||
| 40406 | 774  | 
  \item \hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}} also works on termination goals,
 | 
| 33858 | 775  | 
using a variation of the size-change principle, together with a  | 
776  | 
  graph decomposition technique (see \cite{krauss_phd} for details).
 | 
|
777  | 
  Three kinds of orders are used internally: \isa{max}, \isa{min},
 | 
|
778  | 
  and \isa{ms} (multiset), which is only available when the theory
 | 
|
779  | 
  \isa{Multiset} is loaded. When no order kinds are given, they are
 | 
|
780  | 
tried in order. The search for a termination proof uses SAT solving  | 
|
781  | 
internally.  | 
|
782  | 
||
| 42930 | 783  | 
  For local descent proofs, the \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}} modifiers are
 | 
784  | 
  accepted (as for \hyperlink{method.auto}{\mbox{\isa{auto}}}).
 | 
|
| 33858 | 785  | 
|
| 
45944
 
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
 
bulwahn 
parents: 
45943 
diff
changeset
 | 
786  | 
  \item \hyperlink{method.HOL.induction-schema}{\mbox{\isa{induction{\isaliteral{5F}{\isacharunderscore}}schema}}} derives user-specified
 | 
| 46283 | 787  | 
induction rules from well-founded induction and completeness of  | 
788  | 
patterns. This factors out some operations that are done internally  | 
|
789  | 
by the function package and makes them available separately. See  | 
|
790  | 
\verb|~~/src/HOL/ex/Induction_Schema.thy| for examples.  | 
|
| 
45944
 
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
 
bulwahn 
parents: 
45943 
diff
changeset
 | 
791  | 
|
| 28788 | 792  | 
  \end{description}%
 | 
| 26849 | 793  | 
\end{isamarkuptext}%
 | 
794  | 
\isamarkuptrue%  | 
|
795  | 
%  | 
|
| 40171 | 796  | 
\isamarkupsubsection{Functions with explicit partiality%
 | 
797  | 
}  | 
|
798  | 
\isamarkuptrue%  | 
|
799  | 
%  | 
|
800  | 
\begin{isamarkuptext}%
 | 
|
801  | 
\begin{matharray}{rcl}
 | 
|
| 40406 | 802  | 
    \indexdef{HOL}{command}{partial\_function}\hypertarget{command.HOL.partial-function}{\hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
803  | 
    \indexdef{HOL}{attribute}{partial\_function\_mono}\hypertarget{attribute.HOL.partial-function-mono}{\hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}}} & : & \isa{attribute} \\
 | 
|
| 40171 | 804  | 
  \end{matharray}
 | 
805  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
806  | 
  \begin{railoutput}
 | 
| 42662 | 807  | 
\rail@begin{5}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
808  | 
\rail@term{\hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
809  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
810  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
811  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
812  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
813  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 42617 | 814  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
815  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
816  | 
\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
817  | 
\rail@cr{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
818  | 
\rail@term{\isa{\isakeyword{where}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
819  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
820  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
821  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
822  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
823  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
824  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
825  | 
\end{railoutput}
 | 
| 42617 | 826  | 
|
| 40171 | 827  | 
|
828  | 
  \begin{description}
 | 
|
829  | 
||
| 42617 | 830  | 
  \item \hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}mode{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} defines
 | 
831  | 
recursive functions based on fixpoints in complete partial  | 
|
832  | 
orders. No termination proof is required from the user or  | 
|
833  | 
constructed internally. Instead, the possibility of non-termination  | 
|
834  | 
is modelled explicitly in the result type, which contains an  | 
|
835  | 
explicit bottom element.  | 
|
| 40171 | 836  | 
|
837  | 
Pattern matching and mutual recursion are currently not supported.  | 
|
838  | 
Thus, the specification consists of a single function described by a  | 
|
839  | 
single recursive equation.  | 
|
840  | 
||
841  | 
There are no fixed syntactic restrictions on the body of the  | 
|
842  | 
function, but the induced functional must be provably monotonic  | 
|
843  | 
wrt.\ the underlying order. The monotonicitity proof is performed  | 
|
844  | 
internally, and the definition is rejected when it fails. The proof  | 
|
845  | 
can be influenced by declaring hints using the  | 
|
| 40406 | 846  | 
  \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}} attribute.
 | 
| 40171 | 847  | 
|
848  | 
  The mandatory \isa{mode} argument specifies the mode of operation
 | 
|
849  | 
of the command, which directly corresponds to a complete partial  | 
|
850  | 
order on the result type. By default, the following modes are  | 
|
| 42123 | 851  | 
defined:  | 
| 40171 | 852  | 
|
853  | 
  \begin{description}
 | 
|
| 46283 | 854  | 
|
| 40171 | 855  | 
  \item \isa{option} defines functions that map into the \isa{option} type. Here, the value \isa{None} is used to model a
 | 
| 46283 | 856  | 
  non-terminating computation. Monotonicity requires that if \isa{None} is returned by a recursive call, then the overall result must
 | 
857  | 
  also be \isa{None}. This is best achieved through the use of the
 | 
|
858  | 
  monadic operator \isa{{\isaliteral{22}{\isachardoublequote}}Option{\isaliteral{2E}{\isachardot}}bind{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 42123 | 859  | 
|
| 40171 | 860  | 
  \item \isa{tailrec} defines functions with an arbitrary result
 | 
| 40406 | 861  | 
  type and uses the slightly degenerated partial order where \isa{{\isaliteral{22}{\isachardoublequote}}undefined{\isaliteral{22}{\isachardoublequote}}} is the bottom element.  Now, monotonicity requires that
 | 
| 40171 | 862  | 
  if \isa{undefined} is returned by a recursive call, then the
 | 
863  | 
  overall result must also be \isa{undefined}. In practice, this is
 | 
|
864  | 
only satisfied when each recursive call is a tail call, whose result  | 
|
865  | 
is directly returned. Thus, this mode of operation allows the  | 
|
866  | 
definition of arbitrary tail-recursive functions.  | 
|
| 46283 | 867  | 
|
| 40171 | 868  | 
  \end{description}
 | 
869  | 
||
870  | 
Experienced users may define new modes by instantiating the locale  | 
|
| 40406 | 871  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}definitions{\isaliteral{22}{\isachardoublequote}}} appropriately.
 | 
| 40171 | 872  | 
|
| 40406 | 873  | 
  \item \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}} declares rules for
 | 
| 40171 | 874  | 
use in the internal monononicity proofs of partial function  | 
875  | 
definitions.  | 
|
876  | 
||
877  | 
  \end{description}%
 | 
|
878  | 
\end{isamarkuptext}%
 | 
|
879  | 
\isamarkuptrue%  | 
|
880  | 
%  | 
|
| 26849 | 881  | 
\isamarkupsubsection{Old-style recursive function definitions (TFL)%
 | 
882  | 
}  | 
|
883  | 
\isamarkuptrue%  | 
|
884  | 
%  | 
|
885  | 
\begin{isamarkuptext}%
 | 
|
| 46280 | 886  | 
\begin{matharray}{rcl}
 | 
| 40406 | 887  | 
    \indexdef{HOL}{command}{recdef}\hypertarget{command.HOL.recdef}{\hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
888  | 
    \indexdef{HOL}{command}{recdef\_tc}\hypertarget{command.HOL.recdef-tc}{\hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 26849 | 889  | 
  \end{matharray}
 | 
890  | 
||
| 46280 | 891  | 
  The old TFL commands \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} and \hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}} for defining recursive are mostly obsolete; \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} or \hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}} should be used instead.
 | 
892  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
893  | 
  \begin{railoutput}
 | 
| 42662 | 894  | 
\rail@begin{5}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
895  | 
\rail@term{\hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
896  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
897  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
898  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
899  | 
\rail@term{\isa{\isakeyword{permissive}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
900  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
901  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
902  | 
\rail@cr{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
903  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
904  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
905  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
906  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
907  | 
\rail@nextplus{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
908  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
909  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
910  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
911  | 
\rail@nont{\isa{hints}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
912  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
913  | 
\rail@end  | 
| 42662 | 914  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
915  | 
\rail@nont{\isa{recdeftc}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
916  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
917  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
918  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
919  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
920  | 
\rail@nont{\isa{tc}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
921  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
922  | 
\rail@begin{2}{\isa{hints}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
923  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
924  | 
\rail@term{\isa{\isakeyword{hints}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
925  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
926  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
927  | 
\rail@cnont{\isa{recdefmod}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
928  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
929  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
930  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
931  | 
\rail@begin{4}{\isa{recdefmod}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
932  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
933  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
934  | 
\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
935  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
936  | 
\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
937  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
938  | 
\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
939  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
940  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
941  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
942  | 
\rail@term{\isa{add}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
943  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
944  | 
\rail@term{\isa{del}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
945  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
946  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
947  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
948  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
949  | 
\rail@nont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
950  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
951  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
952  | 
\rail@begin{2}{\isa{tc}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
953  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
954  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
955  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
956  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
957  | 
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
958  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
959  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
960  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
961  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
962  | 
|
| 26849 | 963  | 
|
| 28788 | 964  | 
  \begin{description}
 | 
| 42123 | 965  | 
|
| 28788 | 966  | 
  \item \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} defines general well-founded
 | 
| 26849 | 967  | 
recursive functions (using the TFL package), see also  | 
| 40406 | 968  | 
  \cite{isabelle-HOL}.  The ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}permissive{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' option tells
 | 
| 26849 | 969  | 
TFL to recover from failed proof attempts, returning unfinished  | 
| 40406 | 970  | 
  results.  The \isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}, \isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}, and \isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf} hints refer to auxiliary rules to be used in the internal
 | 
| 26902 | 971  | 
  automated proof process of TFL.  Additional \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}
 | 
| 42930 | 972  | 
declarations may be given to tune the context of the Simplifier  | 
973  | 
  (cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\
 | 
|
974  | 
  \secref{sec:classical}).
 | 
|
| 42123 | 975  | 
|
| 40406 | 976  | 
  \item \hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{28}{\isacharparenleft}}i{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} recommences the
 | 
| 26849 | 977  | 
  proof for leftover termination condition number \isa{i} (default
 | 
| 26902 | 978  | 
  1) as generated by a \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} definition of
 | 
| 26849 | 979  | 
  constant \isa{c}.
 | 
| 42123 | 980  | 
|
| 26902 | 981  | 
  Note that in most cases, \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} is able to finish
 | 
| 26849 | 982  | 
its internal proofs without manual intervention.  | 
983  | 
||
| 28788 | 984  | 
  \end{description}
 | 
| 26849 | 985  | 
|
| 26902 | 986  | 
  \medskip Hints for \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} may be also declared
 | 
| 26849 | 987  | 
globally, using the following attributes.  | 
988  | 
||
989  | 
  \begin{matharray}{rcl}
 | 
|
| 40406 | 990  | 
    \indexdef{HOL}{attribute}{recdef\_simp}\hypertarget{attribute.HOL.recdef-simp}{\hyperlink{attribute.HOL.recdef-simp}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}}} & : & \isa{attribute} \\
 | 
991  | 
    \indexdef{HOL}{attribute}{recdef\_cong}\hypertarget{attribute.HOL.recdef-cong}{\hyperlink{attribute.HOL.recdef-cong}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}}} & : & \isa{attribute} \\
 | 
|
992  | 
    \indexdef{HOL}{attribute}{recdef\_wf}\hypertarget{attribute.HOL.recdef-wf}{\hyperlink{attribute.HOL.recdef-wf}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}}} & : & \isa{attribute} \\
 | 
|
| 26849 | 993  | 
  \end{matharray}
 | 
994  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
995  | 
  \begin{railoutput}
 | 
| 42662 | 996  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
997  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
998  | 
\rail@term{\hyperlink{attribute.HOL.recdef-simp}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
999  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1000  | 
\rail@term{\hyperlink{attribute.HOL.recdef-cong}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1001  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1002  | 
\rail@term{\hyperlink{attribute.HOL.recdef-wf}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1003  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1004  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1005  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1006  | 
\rail@term{\isa{add}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1007  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1008  | 
\rail@term{\isa{del}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1009  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1010  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1011  | 
\end{railoutput}%
 | 
| 26849 | 1012  | 
\end{isamarkuptext}%
 | 
1013  | 
\isamarkuptrue%  | 
|
1014  | 
%  | 
|
| 42908 | 1015  | 
\isamarkupsection{Datatypes \label{sec:hol-datatype}%
 | 
| 26849 | 1016  | 
}  | 
1017  | 
\isamarkuptrue%  | 
|
1018  | 
%  | 
|
1019  | 
\begin{isamarkuptext}%
 | 
|
| 42908 | 1020  | 
\begin{matharray}{rcl}
 | 
1021  | 
    \indexdef{HOL}{command}{datatype}\hypertarget{command.HOL.datatype}{\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1022  | 
    \indexdef{HOL}{command}{rep\_datatype}\hypertarget{command.HOL.rep-datatype}{\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 26849 | 1023  | 
  \end{matharray}
 | 
1024  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1025  | 
  \begin{railoutput}
 | 
| 42908 | 1026  | 
\rail@begin{2}{}
 | 
1027  | 
\rail@term{\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}}}[]
 | 
|
1028  | 
\rail@plus  | 
|
1029  | 
\rail@nont{\isa{spec}}[]
 | 
|
1030  | 
\rail@nextplus{1}
 | 
|
1031  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
|
1032  | 
\rail@endplus  | 
|
1033  | 
\rail@end  | 
|
1034  | 
\rail@begin{3}{}
 | 
|
1035  | 
\rail@term{\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}}[]
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1036  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1037  | 
\rail@nextbar{1}
 | 
| 42908 | 1038  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
1039  | 
\rail@plus  | 
|
1040  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
1041  | 
\rail@nextplus{2}
 | 
|
1042  | 
\rail@endplus  | 
|
1043  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1044  | 
\rail@endbar  | 
| 42908 | 1045  | 
\rail@plus  | 
1046  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
|
1047  | 
\rail@nextplus{1}
 | 
|
1048  | 
\rail@endplus  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1049  | 
\rail@end  | 
| 42908 | 1050  | 
\rail@begin{2}{\isa{spec}}
 | 
| 
45839
 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 
wenzelm 
parents: 
45768 
diff
changeset
 | 
1051  | 
\rail@nont{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}[]
 | 
| 42908 | 1052  | 
\rail@bar  | 
1053  | 
\rail@nextbar{1}
 | 
|
1054  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
|
1055  | 
\rail@endbar  | 
|
1056  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
|
1057  | 
\rail@plus  | 
|
1058  | 
\rail@nont{\isa{cons}}[]
 | 
|
1059  | 
\rail@nextplus{1}
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1060  | 
\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1061  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1062  | 
\rail@end  | 
| 42908 | 1063  | 
\rail@begin{2}{\isa{cons}}
 | 
1064  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
1065  | 
\rail@plus  | 
|
1066  | 
\rail@nextplus{1}
 | 
|
1067  | 
\rail@cnont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
|
1068  | 
\rail@endplus  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1069  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1070  | 
\rail@nextbar{1}
 | 
| 42908 | 1071  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1072  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1073  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1074  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1075  | 
|
| 26849 | 1076  | 
|
| 28788 | 1077  | 
  \begin{description}
 | 
| 26849 | 1078  | 
|
| 42908 | 1079  | 
  \item \hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}} defines inductive datatypes in
 | 
1080  | 
HOL.  | 
|
1081  | 
||
1082  | 
  \item \hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}} represents existing types as
 | 
|
| 42909 | 1083  | 
datatypes.  | 
1084  | 
||
1085  | 
  For foundational reasons, some basic types such as \isa{nat}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2B}{\isacharplus}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{22}{\isachardoublequote}}}, \isa{bool} and \isa{unit} are
 | 
|
1086  | 
  introduced by more primitive means using \indexref{}{command}{typedef}\hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}}.  To
 | 
|
1087  | 
  recover the rich infrastructure of \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}} (e.g.\ rules
 | 
|
1088  | 
  for \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} and the primitive recursion
 | 
|
1089  | 
combinators), such types may be represented as actual datatypes  | 
|
1090  | 
later. This is done by specifying the constructors of the desired  | 
|
1091  | 
type, and giving a proof of the induction rule, distinctness and  | 
|
1092  | 
injectivity of constructors.  | 
|
1093  | 
||
1094  | 
For example, see \verb|~~/src/HOL/Sum_Type.thy| for the  | 
|
1095  | 
representation of the primitive sum type as fully-featured datatype.  | 
|
| 42908 | 1096  | 
|
1097  | 
  \end{description}
 | 
|
1098  | 
||
| 42909 | 1099  | 
  The generated rules for \hyperlink{method.induct}{\mbox{\isa{induct}}} and \hyperlink{method.cases}{\mbox{\isa{cases}}} provide
 | 
1100  | 
case names according to the given constructors, while parameters are  | 
|
1101  | 
  named after the types (see also \secref{sec:cases-induct}).
 | 
|
| 42908 | 1102  | 
|
1103  | 
  See \cite{isabelle-HOL} for more details on datatypes, but beware of
 | 
|
1104  | 
the old-style theory syntax being used there! Apart from proper  | 
|
1105  | 
proof methods for case-analysis and induction, there are also  | 
|
1106  | 
  emulations of ML tactics \hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}} and \hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}} available, see \secref{sec:hol-induct-tac}; these admit
 | 
|
1107  | 
to refer directly to the internal structure of subgoals (including  | 
|
1108  | 
internally bound parameters).%  | 
|
1109  | 
\end{isamarkuptext}%
 | 
|
1110  | 
\isamarkuptrue%  | 
|
1111  | 
%  | 
|
| 
42910
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1112  | 
\isamarkupsubsubsection{Examples%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1113  | 
}  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1114  | 
\isamarkuptrue%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1115  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1116  | 
\begin{isamarkuptext}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1117  | 
We define a type of finite sequences, with slightly different  | 
| 46525 | 1118  | 
  names than the existing \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequote}}} that is already in \isa{Main}:%
 | 
| 
42910
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1119  | 
\end{isamarkuptext}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1120  | 
\isamarkuptrue%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1121  | 
\isacommand{datatype}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1122  | 
\ {\isaliteral{27}{\isacharprime}}a\ seq\ {\isaliteral{3D}{\isacharequal}}\ Empty\ {\isaliteral{7C}{\isacharbar}}\ Seq\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ seq{\isaliteral{22}{\isachardoublequoteclose}}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1123  | 
\begin{isamarkuptext}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1124  | 
We can now prove some simple lemma by structural induction:%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1125  | 
\end{isamarkuptext}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1126  | 
\isamarkuptrue%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1127  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1128  | 
\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ xs\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1129  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1130  | 
\isadelimproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1131  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1132  | 
\endisadelimproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1133  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1134  | 
\isatagproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1135  | 
\isacommand{proof}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1136  | 
\ {\isaliteral{28}{\isacharparenleft}}induct\ xs\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\isanewline
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1137  | 
\ \ \isacommand{case}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1138  | 
\ Empty%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1139  | 
\begin{isamarkuptxt}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1140  | 
This case can be proved using the simplifier: the freeness  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1141  | 
    properties of the datatype are already declared as \hyperlink{attribute.simp}{\mbox{\isa{simp}}} rules.%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1142  | 
\end{isamarkuptxt}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1143  | 
\isamarkuptrue%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1144  | 
\ \ \isacommand{show}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1145  | 
\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ Empty\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Empty{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1146  | 
\ \ \ \ \isacommand{by}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1147  | 
\ simp\isanewline  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1148  | 
\isacommand{next}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1149  | 
\isanewline  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1150  | 
\ \ \isacommand{case}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1151  | 
\ {\isaliteral{28}{\isacharparenleft}}Seq\ y\ ys{\isaliteral{29}{\isacharparenright}}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1152  | 
\begin{isamarkuptxt}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1153  | 
The step case is proved similarly.%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1154  | 
\end{isamarkuptxt}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1155  | 
\isamarkuptrue%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1156  | 
\ \ \isacommand{show}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1157  | 
\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ {\isaliteral{28}{\isacharparenleft}}Seq\ y\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Seq\ y\ ys{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1158  | 
\ \ \ \ \isacommand{using}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1159  | 
\ {\isaliteral{60}{\isacharbackquoteopen}}Seq\ y\ ys\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ ys{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{by}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1160  | 
\ simp\isanewline  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1161  | 
\isacommand{qed}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1162  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1163  | 
\endisatagproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1164  | 
{\isafoldproof}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1165  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1166  | 
\isadelimproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1167  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1168  | 
\endisadelimproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1169  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1170  | 
\begin{isamarkuptext}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1171  | 
Here is a more succinct version of the same proof:%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1172  | 
\end{isamarkuptext}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1173  | 
\isamarkuptrue%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1174  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1175  | 
\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ xs\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1176  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1177  | 
\isadelimproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1178  | 
\ \ %  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1179  | 
\endisadelimproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1180  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1181  | 
\isatagproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1182  | 
\isacommand{by}\isamarkupfalse%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1183  | 
\ {\isaliteral{28}{\isacharparenleft}}induct\ xs\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1184  | 
\endisatagproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1185  | 
{\isafoldproof}%
 | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1186  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1187  | 
\isadelimproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1188  | 
%  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1189  | 
\endisadelimproof  | 
| 
 
6834af822a8b
updated and simplified HOL datatype examples (NB: special treatment of distinctness has been discontinued in the vicinity of 542b34b178ec);
 
wenzelm 
parents: 
42909 
diff
changeset
 | 
1190  | 
%  | 
| 42908 | 1191  | 
\isamarkupsection{Records \label{sec:hol-record}%
 | 
1192  | 
}  | 
|
1193  | 
\isamarkuptrue%  | 
|
1194  | 
%  | 
|
1195  | 
\begin{isamarkuptext}%
 | 
|
1196  | 
In principle, records merely generalize the concept of tuples, where  | 
|
1197  | 
components may be addressed by labels instead of just position. The  | 
|
1198  | 
logical infrastructure of records in Isabelle/HOL is slightly more  | 
|
1199  | 
advanced, though, supporting truly extensible record schemes. This  | 
|
1200  | 
admits operations that are polymorphic with respect to record  | 
|
1201  | 
extension, yielding ``object-oriented'' effects like (single)  | 
|
1202  | 
  inheritance.  See also \cite{NaraschewskiW-TPHOLs98} for more
 | 
|
1203  | 
details on object-oriented verification and record subtyping in HOL.%  | 
|
1204  | 
\end{isamarkuptext}%
 | 
|
1205  | 
\isamarkuptrue%  | 
|
1206  | 
%  | 
|
1207  | 
\isamarkupsubsection{Basic concepts%
 | 
|
1208  | 
}  | 
|
1209  | 
\isamarkuptrue%  | 
|
1210  | 
%  | 
|
1211  | 
\begin{isamarkuptext}%
 | 
|
1212  | 
Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
 | 
|
1213  | 
at the level of terms and types. The notation is as follows:  | 
|
1214  | 
||
1215  | 
  \begin{center}
 | 
|
1216  | 
  \begin{tabular}{l|l|l}
 | 
|
1217  | 
& record terms & record types \\ \hline  | 
|
1218  | 
    fixed & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1219  | 
    schematic & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ m{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} &
 | 
|
1220  | 
      \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ M{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1221  | 
  \end{tabular}
 | 
|
1222  | 
  \end{center}
 | 
|
1223  | 
||
1224  | 
  \noindent The ASCII representation of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{7C}{\isacharbar}}\ x\ {\isaliteral{3D}{\isacharequal}}\ a\ {\isaliteral{7C}{\isacharbar}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
1225  | 
||
1226  | 
  A fixed record \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} has field \isa{x} of value
 | 
|
1227  | 
  \isa{a} and field \isa{y} of value \isa{b}.  The corresponding
 | 
|
1228  | 
  type is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}, assuming that \isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1229  | 
  and \isa{{\isaliteral{22}{\isachardoublequote}}b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
1230  | 
||
1231  | 
  A record scheme like \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ m{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} contains fields
 | 
|
1232  | 
  \isa{x} and \isa{y} as before, but also possibly further fields
 | 
|
1233  | 
  as indicated by the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' notation (which is actually part
 | 
|
1234  | 
  of the syntax).  The improper field ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' of a record
 | 
|
1235  | 
  scheme is called the \emph{more part}.  Logically it is just a free
 | 
|
1236  | 
variable, which is occasionally referred to as ``row variable'' in  | 
|
1237  | 
the literature. The more part of a record scheme may be  | 
|
1238  | 
instantiated by zero or more further components. For example, the  | 
|
1239  | 
  previous scheme may get instantiated to \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ z\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ m{\isaliteral{27}{\isacharprime}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}, where \isa{m{\isaliteral{27}{\isacharprime}}} refers to a different more part.
 | 
|
1240  | 
Fixed records are special instances of record schemes, where  | 
|
1241  | 
  ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' is properly terminated by the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ unit{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1242  | 
  element.  In fact, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} is just an abbreviation
 | 
|
1243  | 
  for \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 26849 | 1244  | 
|
| 42908 | 1245  | 
\medskip Two key observations make extensible records in a simply  | 
1246  | 
typed language like HOL work out:  | 
|
1247  | 
||
1248  | 
  \begin{enumerate}
 | 
|
1249  | 
||
1250  | 
\item the more part is internalized, as a free term or type  | 
|
1251  | 
variable,  | 
|
1252  | 
||
1253  | 
\item field names are externalized, they cannot be accessed within  | 
|
1254  | 
the logic as first-class values.  | 
|
1255  | 
||
1256  | 
  \end{enumerate}
 | 
|
1257  | 
||
1258  | 
\medskip In Isabelle/HOL record types have to be defined explicitly,  | 
|
1259  | 
fixing their field names and types, and their (optional) parent  | 
|
1260  | 
record. Afterwards, records may be formed using above syntax, while  | 
|
1261  | 
obeying the canonical order of fields as given by their declaration.  | 
|
1262  | 
The record package provides several standard operations like  | 
|
1263  | 
selectors and updates. The common setup for various generic proof  | 
|
1264  | 
tools enable succinct reasoning patterns. See also the Isabelle/HOL  | 
|
1265  | 
  tutorial \cite{isabelle-hol-book} for further instructions on using
 | 
|
1266  | 
records in practice.%  | 
|
1267  | 
\end{isamarkuptext}%
 | 
|
1268  | 
\isamarkuptrue%  | 
|
1269  | 
%  | 
|
1270  | 
\isamarkupsubsection{Record specifications%
 | 
|
1271  | 
}  | 
|
1272  | 
\isamarkuptrue%  | 
|
1273  | 
%  | 
|
1274  | 
\begin{isamarkuptext}%
 | 
|
1275  | 
\begin{matharray}{rcl}
 | 
|
1276  | 
    \indexdef{HOL}{command}{record}\hypertarget{command.HOL.record}{\hyperlink{command.HOL.record}{\mbox{\isa{\isacommand{record}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1277  | 
  \end{matharray}
 | 
|
| 26849 | 1278  | 
|
| 42908 | 1279  | 
  \begin{railoutput}
 | 
1280  | 
\rail@begin{4}{}
 | 
|
1281  | 
\rail@term{\hyperlink{command.HOL.record}{\mbox{\isa{\isacommand{record}}}}}[]
 | 
|
1282  | 
\rail@nont{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}[]
 | 
|
1283  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
|
1284  | 
\rail@cr{2}
 | 
|
1285  | 
\rail@bar  | 
|
1286  | 
\rail@nextbar{3}
 | 
|
1287  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
|
1288  | 
\rail@term{\isa{{\isaliteral{2B}{\isacharplus}}}}[]
 | 
|
1289  | 
\rail@endbar  | 
|
1290  | 
\rail@plus  | 
|
| 
46494
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1291  | 
\rail@nont{\isa{constdecl}}[]
 | 
| 42908 | 1292  | 
\rail@nextplus{3}
 | 
1293  | 
\rail@endplus  | 
|
1294  | 
\rail@end  | 
|
| 
46494
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1295  | 
\rail@begin{2}{\isa{constdecl}}
 | 
| 
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1296  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1297  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1298  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
| 
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1299  | 
\rail@bar  | 
| 
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1300  | 
\rail@nextbar{1}
 | 
| 
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1301  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
| 
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1302  | 
\rail@endbar  | 
| 
 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 
wenzelm 
parents: 
46457 
diff
changeset
 | 
1303  | 
\rail@end  | 
| 42908 | 1304  | 
\end{railoutput}
 | 
1305  | 
||
1306  | 
||
1307  | 
  \begin{description}
 | 
|
1308  | 
||
1309  | 
  \item \hyperlink{command.HOL.record}{\mbox{\isa{\isacommand{record}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\ {\isaliteral{2B}{\isacharplus}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} defines extensible record type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}},
 | 
|
1310  | 
  derived from the optional parent record \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} by adding new
 | 
|
1311  | 
  field components \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} etc.
 | 
|
1312  | 
||
1313  | 
  The type variables of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} need to be
 | 
|
1314  | 
  covered by the (distinct) parameters \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}}.  Type constructor \isa{t} has to be new, while \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} needs to specify an instance of an existing record type.  At
 | 
|
1315  | 
  least one new field \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} has to be specified.
 | 
|
1316  | 
Basically, field names need to belong to a unique record. This is  | 
|
1317  | 
not a real restriction in practice, since fields are qualified by  | 
|
1318  | 
the record name internally.  | 
|
1319  | 
||
1320  | 
  The parent record specification \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} is optional; if omitted
 | 
|
1321  | 
  \isa{t} becomes a root record.  The hierarchy of all records
 | 
|
1322  | 
declared within a theory context forms a forest structure, i.e.\ a  | 
|
1323  | 
set of trees starting with a root record each. There is no way to  | 
|
1324  | 
merge multiple parent records!  | 
|
1325  | 
||
1326  | 
  For convenience, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} is made a
 | 
|
1327  | 
  type abbreviation for the fixed record type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}, likewise is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{5F}{\isacharunderscore}}scheme{\isaliteral{22}{\isachardoublequote}}} made an abbreviation for
 | 
|
1328  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 26849 | 1329  | 
|
| 28788 | 1330  | 
  \end{description}%
 | 
| 26849 | 1331  | 
\end{isamarkuptext}%
 | 
1332  | 
\isamarkuptrue%  | 
|
1333  | 
%  | 
|
| 42908 | 1334  | 
\isamarkupsubsection{Record operations%
 | 
1335  | 
}  | 
|
1336  | 
\isamarkuptrue%  | 
|
1337  | 
%  | 
|
1338  | 
\begin{isamarkuptext}%
 | 
|
1339  | 
Any record definition of the form presented above produces certain  | 
|
1340  | 
standard operations. Selectors and updates are provided for any  | 
|
1341  | 
  field, including the improper one ``\isa{more}''.  There are also
 | 
|
1342  | 
cumulative record constructor functions. To simplify the  | 
|
1343  | 
  presentation below, we assume for now that \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} is a root record with fields \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
1344  | 
||
1345  | 
  \medskip \textbf{Selectors} and \textbf{updates} are available for
 | 
|
1346  | 
  any field (including ``\isa{more}''):
 | 
|
1347  | 
||
1348  | 
  \begin{matharray}{lll}
 | 
|
1349  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1350  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{5F}{\isacharunderscore}}update{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1351  | 
  \end{matharray}
 | 
|
1352  | 
||
1353  | 
  There is special syntax for application of updates: \isa{{\isaliteral{22}{\isachardoublequote}}r{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} abbreviates term \isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{5F}{\isacharunderscore}}update\ a\ r{\isaliteral{22}{\isachardoublequote}}}.  Further notation for
 | 
|
1354  | 
  repeated updates is also available: \isa{{\isaliteral{22}{\isachardoublequote}}r{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} may be written \isa{{\isaliteral{22}{\isachardoublequote}}r{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}.  Note that
 | 
|
1355  | 
because of postfix notation the order of fields shown here is  | 
|
1356  | 
reverse than in the actual term. Since repeated updates are just  | 
|
1357  | 
  function applications, fields may be freely permuted in \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}, as far as logical equality is concerned.
 | 
|
1358  | 
Thus commutativity of independent updates can be proven within the  | 
|
1359  | 
logic for any two fields, but not as a general theorem.  | 
|
1360  | 
||
1361  | 
  \medskip The \textbf{make} operation provides a cumulative record
 | 
|
1362  | 
constructor function:  | 
|
1363  | 
||
1364  | 
  \begin{matharray}{lll}
 | 
|
1365  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}make{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1366  | 
  \end{matharray}
 | 
|
1367  | 
||
1368  | 
\medskip We now reconsider the case of non-root records, which are  | 
|
1369  | 
derived of some parent. In general, the latter may depend on  | 
|
1370  | 
  another parent as well, resulting in a list of \emph{ancestor
 | 
|
1371  | 
records}. Appending the lists of fields of all ancestors results in  | 
|
1372  | 
a certain field prefix. The record package automatically takes care  | 
|
1373  | 
of this by lifting operations over this context of ancestor fields.  | 
|
1374  | 
  Assuming that \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} has ancestor
 | 
|
1375  | 
  fields \isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C72686F3E}{\isasymrho}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub k\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C72686F3E}{\isasymrho}}\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{22}{\isachardoublequote}}},
 | 
|
1376  | 
the above record operations will get the following types:  | 
|
1377  | 
||
1378  | 
\medskip  | 
|
1379  | 
  \begin{tabular}{lll}
 | 
|
1380  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1381  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{5F}{\isacharunderscore}}update{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1382  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}make{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C72686F3E}{\isasymrho}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C72686F3E}{\isasymrho}}\isaliteral{5C3C5E7375623E}{}\isactrlsub k\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1383  | 
  \end{tabular}
 | 
|
1384  | 
\medskip  | 
|
1385  | 
||
1386  | 
\noindent Some further operations address the extension aspect of a  | 
|
1387  | 
  derived record scheme specifically: \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}fields{\isaliteral{22}{\isachardoublequote}}} produces a
 | 
|
1388  | 
record fragment consisting of exactly the new fields introduced here  | 
|
1389  | 
  (the result may serve as a more part elsewhere); \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}extend{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1390  | 
  takes a fixed record and adds a given more part; \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}truncate{\isaliteral{22}{\isachardoublequote}}} restricts a record scheme to a fixed record.
 | 
|
1391  | 
||
1392  | 
\medskip  | 
|
1393  | 
  \begin{tabular}{lll}
 | 
|
1394  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}fields{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1395  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}extend{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1396  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}truncate{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1397  | 
  \end{tabular}
 | 
|
1398  | 
\medskip  | 
|
1399  | 
||
1400  | 
  \noindent Note that \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}make{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}fields{\isaliteral{22}{\isachardoublequote}}} coincide
 | 
|
1401  | 
for root records.%  | 
|
1402  | 
\end{isamarkuptext}%
 | 
|
1403  | 
\isamarkuptrue%  | 
|
1404  | 
%  | 
|
1405  | 
\isamarkupsubsection{Derived rules and proof tools%
 | 
|
| 26849 | 1406  | 
}  | 
1407  | 
\isamarkuptrue%  | 
|
1408  | 
%  | 
|
1409  | 
\begin{isamarkuptext}%
 | 
|
| 42908 | 1410  | 
The record package proves several results internally, declaring  | 
1411  | 
these facts to appropriate proof tools. This enables users to  | 
|
1412  | 
reason about record structures quite conveniently. Assume that  | 
|
1413  | 
  \isa{t} is a record type as specified above.
 | 
|
1414  | 
||
1415  | 
  \begin{enumerate}
 | 
|
1416  | 
||
1417  | 
\item Standard conversions for selectors or updates applied to  | 
|
1418  | 
record constructor terms are made part of the default Simplifier  | 
|
1419  | 
context; thus proofs by reduction of basic operations merely require  | 
|
1420  | 
  the \hyperlink{method.simp}{\mbox{\isa{simp}}} method without further arguments.  These rules
 | 
|
1421  | 
  are available as \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}simps{\isaliteral{22}{\isachardoublequote}}}, too.
 | 
|
1422  | 
||
1423  | 
\item Selectors applied to updated records are automatically reduced  | 
|
1424  | 
by an internal simplification procedure, which is also part of the  | 
|
1425  | 
standard Simplifier setup.  | 
|
1426  | 
||
1427  | 
  \item Inject equations of a form analogous to \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{2C}{\isacharcomma}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{27}{\isacharprime}}{\isaliteral{2C}{\isacharcomma}}\ y{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}} are declared to the Simplifier and Classical
 | 
|
1428  | 
  Reasoner as \hyperlink{attribute.iff}{\mbox{\isa{iff}}} rules.  These rules are available as
 | 
|
1429  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}iffs{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
1430  | 
||
1431  | 
  \item The introduction rule for record equality analogous to \isa{{\isaliteral{22}{\isachardoublequote}}x\ r\ {\isaliteral{3D}{\isacharequal}}\ x\ r{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ r\ {\isaliteral{3D}{\isacharequal}}\ y\ r{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ r\ {\isaliteral{3D}{\isacharequal}}\ r{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}} is declared to the Simplifier,
 | 
|
1432  | 
  and as the basic rule context as ``\hyperlink{attribute.intro}{\mbox{\isa{intro}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}''.
 | 
|
1433  | 
  The rule is called \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}equality{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
1434  | 
||
1435  | 
\item Representations of arbitrary record expressions as canonical  | 
|
1436  | 
  constructor terms are provided both in \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} format (cf.\ the generic proof methods of the same name,
 | 
|
1437  | 
  \secref{sec:cases-induct}).  Several variations are available, for
 | 
|
1438  | 
fixed records, record schemes, more parts etc.  | 
|
1439  | 
||
1440  | 
The generic proof methods are sufficiently smart to pick the most  | 
|
1441  | 
sensible rule according to the type of the indicated record  | 
|
1442  | 
  expression: users just need to apply something like ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}cases\ r{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' to a certain proof problem.
 | 
|
1443  | 
||
1444  | 
  \item The derived record operations \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}make{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}fields{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}extend{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}truncate{\isaliteral{22}{\isachardoublequote}}} are \emph{not}
 | 
|
1445  | 
treated automatically, but usually need to be expanded by hand,  | 
|
1446  | 
  using the collective fact \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}defs{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
1447  | 
||
1448  | 
  \end{enumerate}%
 | 
|
1449  | 
\end{isamarkuptext}%
 | 
|
1450  | 
\isamarkuptrue%  | 
|
1451  | 
%  | 
|
| 42911 | 1452  | 
\isamarkupsubsubsection{Examples%
 | 
1453  | 
}  | 
|
1454  | 
\isamarkuptrue%  | 
|
1455  | 
%  | 
|
1456  | 
\begin{isamarkuptext}%
 | 
|
1457  | 
See \verb|~~/src/HOL/ex/Records.thy|, for example.%  | 
|
1458  | 
\end{isamarkuptext}%
 | 
|
1459  | 
\isamarkuptrue%  | 
|
1460  | 
%  | 
|
| 42908 | 1461  | 
\isamarkupsection{Adhoc tuples%
 | 
1462  | 
}  | 
|
1463  | 
\isamarkuptrue%  | 
|
1464  | 
%  | 
|
1465  | 
\begin{isamarkuptext}%
 | 
|
1466  | 
\begin{matharray}{rcl}
 | 
|
1467  | 
    \indexdef{HOL}{attribute}{split\_format}\hypertarget{attribute.HOL.split-format}{\hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{attribute} \\
 | 
|
1468  | 
  \end{matharray}
 | 
|
1469  | 
||
1470  | 
  \begin{railoutput}
 | 
|
1471  | 
\rail@begin{2}{}
 | 
|
1472  | 
\rail@term{\hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}}[]
 | 
|
1473  | 
\rail@bar  | 
|
1474  | 
\rail@nextbar{1}
 | 
|
1475  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
|
1476  | 
\rail@term{\isa{complete}}[]
 | 
|
1477  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
|
1478  | 
\rail@endbar  | 
|
1479  | 
\rail@end  | 
|
1480  | 
\end{railoutput}
 | 
|
1481  | 
||
| 26849 | 1482  | 
|
1483  | 
  \begin{description}
 | 
|
1484  | 
||
| 42908 | 1485  | 
  \item \hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}\ \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}complete{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} causes
 | 
1486  | 
arguments in function applications to be represented canonically  | 
|
1487  | 
according to their tuple type structure.  | 
|
| 26849 | 1488  | 
|
| 42908 | 1489  | 
Note that this operation tends to invent funny names for new local  | 
1490  | 
parameters introduced.  | 
|
| 26849 | 1491  | 
|
| 42908 | 1492  | 
  \end{description}%
 | 
| 26849 | 1493  | 
\end{isamarkuptext}%
 | 
1494  | 
\isamarkuptrue%  | 
|
1495  | 
%  | 
|
| 42908 | 1496  | 
\isamarkupsection{Typedef axiomatization \label{sec:hol-typedef}%
 | 
| 26849 | 1497  | 
}  | 
1498  | 
\isamarkuptrue%  | 
|
1499  | 
%  | 
|
1500  | 
\begin{isamarkuptext}%
 | 
|
| 46280 | 1501  | 
\begin{matharray}{rcl}
 | 
1502  | 
    \indexdef{HOL}{command}{typedef}\hypertarget{command.HOL.typedef}{\hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1503  | 
  \end{matharray}
 | 
|
1504  | 
||
1505  | 
A Gordon/HOL-style type definition is a certain axiom scheme that  | 
|
1506  | 
identifies a new type with a subset of an existing type. More  | 
|
| 42908 | 1507  | 
precisely, the new type is defined by exhibiting an existing type  | 
1508  | 
  \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}}, a set \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\ set{\isaliteral{22}{\isachardoublequote}}}, and a theorem that proves
 | 
|
1509  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequote}}}.  Thus \isa{A} is a non-empty subset of \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}}, and the new type denotes this subset.  New functions are
 | 
|
1510  | 
postulated that establish an isomorphism between the new type and  | 
|
1511  | 
  the subset.  In general, the type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} may involve type
 | 
|
1512  | 
  variables \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} which means that the type definition
 | 
|
1513  | 
  produces a type constructor \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} depending on
 | 
|
1514  | 
those type arguments.  | 
|
1515  | 
||
1516  | 
The axiomatization can be considered a ``definition'' in the sense  | 
|
1517  | 
of the particular set-theoretic interpretation of HOL  | 
|
1518  | 
  \cite{pitts93}, where the universe of types is required to be
 | 
|
1519  | 
downwards-closed wrt.\ arbitrary non-empty subsets. Thus genuinely  | 
|
1520  | 
  new types introduced by \hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}} stay within the range
 | 
|
1521  | 
  of HOL models by construction.  Note that \indexref{}{command}{type\_synonym}\hyperlink{command.type-synonym}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}}}} from Isabelle/Pure merely introduces syntactic
 | 
|
1522  | 
abbreviations, without any logical significance.  | 
|
| 47859 | 1523  | 
|
| 42908 | 1524  | 
  \begin{railoutput}
 | 
1525  | 
\rail@begin{2}{}
 | 
|
1526  | 
\rail@term{\hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}}[]
 | 
|
1527  | 
\rail@bar  | 
|
1528  | 
\rail@nextbar{1}
 | 
|
1529  | 
\rail@nont{\isa{alt{\isaliteral{5F}{\isacharunderscore}}name}}[]
 | 
|
1530  | 
\rail@endbar  | 
|
1531  | 
\rail@nont{\isa{abs{\isaliteral{5F}{\isacharunderscore}}type}}[]
 | 
|
1532  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
|
1533  | 
\rail@nont{\isa{rep{\isaliteral{5F}{\isacharunderscore}}set}}[]
 | 
|
1534  | 
\rail@end  | 
|
1535  | 
\rail@begin{3}{\isa{alt{\isaliteral{5F}{\isacharunderscore}}name}}
 | 
|
1536  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
|
1537  | 
\rail@bar  | 
|
1538  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
1539  | 
\rail@nextbar{1}
 | 
|
1540  | 
\rail@term{\isa{\isakeyword{open}}}[]
 | 
|
1541  | 
\rail@nextbar{2}
 | 
|
1542  | 
\rail@term{\isa{\isakeyword{open}}}[]
 | 
|
1543  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
1544  | 
\rail@endbar  | 
|
1545  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
|
1546  | 
\rail@end  | 
|
1547  | 
\rail@begin{2}{\isa{abs{\isaliteral{5F}{\isacharunderscore}}type}}
 | 
|
1548  | 
\rail@nont{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}[]
 | 
|
1549  | 
\rail@bar  | 
|
1550  | 
\rail@nextbar{1}
 | 
|
1551  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
|
1552  | 
\rail@endbar  | 
|
1553  | 
\rail@end  | 
|
1554  | 
\rail@begin{2}{\isa{rep{\isaliteral{5F}{\isacharunderscore}}set}}
 | 
|
1555  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
|
1556  | 
\rail@bar  | 
|
1557  | 
\rail@nextbar{1}
 | 
|
1558  | 
\rail@term{\isa{\isakeyword{morphisms}}}[]
 | 
|
1559  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
1560  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
1561  | 
\rail@endbar  | 
|
1562  | 
\rail@end  | 
|
1563  | 
\end{railoutput}
 | 
|
1564  | 
||
1565  | 
||
1566  | 
  \begin{description}
 | 
|
| 26849 | 1567  | 
|
| 42908 | 1568  | 
  \item \hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ A{\isaliteral{22}{\isachardoublequote}}}
 | 
1569  | 
axiomatizes a type definition in the background theory of the  | 
|
1570  | 
current context, depending on a non-emptiness result of the set  | 
|
1571  | 
  \isa{A} that needs to be proven here.  The set \isa{A} may
 | 
|
1572  | 
  contain type variables \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} as specified on the LHS,
 | 
|
1573  | 
but no term variables.  | 
|
1574  | 
||
1575  | 
Even though a local theory specification, the newly introduced type  | 
|
1576  | 
constructor cannot depend on parameters or assumptions of the  | 
|
1577  | 
context: this is structurally impossible in HOL. In contrast, the  | 
|
1578  | 
non-emptiness proof may use local assumptions in unusual situations,  | 
|
1579  | 
which could result in different interpretations in target contexts:  | 
|
1580  | 
  the meaning of the bijection between the representing set \isa{A}
 | 
|
1581  | 
  and the new type \isa{t} may then change in different application
 | 
|
1582  | 
contexts.  | 
|
1583  | 
||
1584  | 
  By default, \hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}} defines both a type
 | 
|
1585  | 
  constructor \isa{t} for the new type, and a term constant \isa{t} for the representing set within the old type.  Use the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}open{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' option to suppress a separate constant definition
 | 
|
1586  | 
  altogether.  The injection from type to set is called \isa{Rep{\isaliteral{5F}{\isacharunderscore}}t},
 | 
|
1587  | 
  its inverse \isa{Abs{\isaliteral{5F}{\isacharunderscore}}t}, unless explicit \hyperlink{keyword.HOL.morphisms}{\mbox{\isa{\isakeyword{morphisms}}}} specification provides alternative names.
 | 
|
1588  | 
||
1589  | 
  The core axiomatization uses the locale predicate \isa{type{\isaliteral{5F}{\isacharunderscore}}definition} as defined in Isabelle/HOL.  Various basic
 | 
|
1590  | 
consequences of that are instantiated accordingly, re-using the  | 
|
1591  | 
locale facts with names derived from the new type constructor. Thus  | 
|
1592  | 
  the generic \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep} is turned into the specific
 | 
|
1593  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}Rep{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{22}{\isachardoublequote}}}, for example.
 | 
|
| 26849 | 1594  | 
|
| 42908 | 1595  | 
  Theorems \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep}, \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}inverse}, and \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}inverse}
 | 
1596  | 
provide the most basic characterization as a corresponding  | 
|
1597  | 
injection/surjection pair (in both directions). The derived rules  | 
|
1598  | 
  \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}inject} and \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}inject} provide a more convenient version of
 | 
|
1599  | 
injectivity, suitable for automated proof tools (e.g.\ in  | 
|
1600  | 
  declarations involving \hyperlink{attribute.simp}{\mbox{\isa{simp}}} or \hyperlink{attribute.iff}{\mbox{\isa{iff}}}).
 | 
|
1601  | 
  Furthermore, the rules \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}cases}~/ \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}induct}, and \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}cases}~/
 | 
|
1602  | 
  \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}induct} provide alternative views on
 | 
|
1603  | 
surjectivity. These rules are already declared as set or type rules  | 
|
1604  | 
  for the generic \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} methods,
 | 
|
1605  | 
respectively.  | 
|
1606  | 
||
1607  | 
An alternative name for the set definition (and other derived  | 
|
1608  | 
entities) may be specified in parentheses; the default is to use  | 
|
1609  | 
  \isa{t} directly.
 | 
|
1610  | 
||
1611  | 
  \end{description}
 | 
|
1612  | 
||
1613  | 
  \begin{warn}
 | 
|
1614  | 
  If you introduce a new type axiomatically, i.e.\ via \indexref{}{command}{typedecl}\hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}} and \indexref{}{command}{axiomatization}\hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}, the minimum requirement
 | 
|
1615  | 
is that it has a non-empty model, to avoid immediate collapse of the  | 
|
1616  | 
HOL logic. Moreover, one needs to demonstrate that the  | 
|
1617  | 
interpretation of such free-form axiomatizations can coexist with  | 
|
1618  | 
  that of the regular \indexdef{}{command}{typedef}\hypertarget{command.typedef}{\hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}}} scheme, and any extension
 | 
|
1619  | 
that other people might have introduced elsewhere (e.g.\ in HOLCF  | 
|
1620  | 
  \cite{MuellerNvOS99}).
 | 
|
1621  | 
  \end{warn}%
 | 
|
1622  | 
\end{isamarkuptext}%
 | 
|
1623  | 
\isamarkuptrue%  | 
|
1624  | 
%  | 
|
1625  | 
\isamarkupsubsubsection{Examples%
 | 
|
1626  | 
}  | 
|
1627  | 
\isamarkuptrue%  | 
|
1628  | 
%  | 
|
1629  | 
\begin{isamarkuptext}%
 | 
|
1630  | 
Type definitions permit the introduction of abstract data  | 
|
1631  | 
types in a safe way, namely by providing models based on already  | 
|
1632  | 
  existing types.  Given some abstract axiomatic description \isa{P}
 | 
|
1633  | 
of a type, this involves two steps:  | 
|
1634  | 
||
1635  | 
  \begin{enumerate}
 | 
|
| 26849 | 1636  | 
|
| 42908 | 1637  | 
  \item Find an appropriate type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} and subset \isa{A} which
 | 
1638  | 
  has the desired properties \isa{P}, and make a type definition
 | 
|
1639  | 
based on this representation.  | 
|
1640  | 
||
1641  | 
  \item Prove that \isa{P} holds for \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} by lifting \isa{P}
 | 
|
1642  | 
from the representation.  | 
|
1643  | 
||
1644  | 
  \end{enumerate}
 | 
|
1645  | 
||
1646  | 
You can later forget about the representation and work solely in  | 
|
1647  | 
  terms of the abstract properties \isa{P}.
 | 
|
| 26849 | 1648  | 
|
| 42908 | 1649  | 
\medskip The following trivial example pulls a three-element type  | 
1650  | 
into existence within the formal logical environment of HOL.%  | 
|
1651  | 
\end{isamarkuptext}%
 | 
|
1652  | 
\isamarkuptrue%  | 
|
1653  | 
\isacommand{typedef}\isamarkupfalse%
 | 
|
1654  | 
\ three\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ False{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}False{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
1655  | 
%  | 
|
1656  | 
\isadelimproof  | 
|
1657  | 
\ \ %  | 
|
1658  | 
\endisadelimproof  | 
|
1659  | 
%  | 
|
1660  | 
\isatagproof  | 
|
1661  | 
\isacommand{by}\isamarkupfalse%
 | 
|
1662  | 
\ blast%  | 
|
1663  | 
\endisatagproof  | 
|
1664  | 
{\isafoldproof}%
 | 
|
1665  | 
%  | 
|
1666  | 
\isadelimproof  | 
|
1667  | 
\isanewline  | 
|
1668  | 
%  | 
|
1669  | 
\endisadelimproof  | 
|
1670  | 
\isanewline  | 
|
1671  | 
\isacommand{definition}\isamarkupfalse%
 | 
|
1672  | 
\ {\isaliteral{22}{\isachardoublequoteopen}}One\ {\isaliteral{3D}{\isacharequal}}\ Abs{\isaliteral{5F}{\isacharunderscore}}three\ {\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
1673  | 
\isacommand{definition}\isamarkupfalse%
 | 
|
1674  | 
\ {\isaliteral{22}{\isachardoublequoteopen}}Two\ {\isaliteral{3D}{\isacharequal}}\ Abs{\isaliteral{5F}{\isacharunderscore}}three\ {\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ False{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
1675  | 
\isacommand{definition}\isamarkupfalse%
 | 
|
1676  | 
\ {\isaliteral{22}{\isachardoublequoteopen}}Three\ {\isaliteral{3D}{\isacharequal}}\ Abs{\isaliteral{5F}{\isacharunderscore}}three\ {\isaliteral{28}{\isacharparenleft}}False{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
1677  | 
\isanewline  | 
|
1678  | 
\isacommand{lemma}\isamarkupfalse%
 | 
|
1679  | 
\ three{\isaliteral{5F}{\isacharunderscore}}distinct{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}One\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Two{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}One\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Three{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}Two\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Three{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
1680  | 
%  | 
|
1681  | 
\isadelimproof  | 
|
1682  | 
\ \ %  | 
|
1683  | 
\endisadelimproof  | 
|
1684  | 
%  | 
|
1685  | 
\isatagproof  | 
|
1686  | 
\isacommand{by}\isamarkupfalse%
 | 
|
1687  | 
\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ One{\isaliteral{5F}{\isacharunderscore}}def\ Two{\isaliteral{5F}{\isacharunderscore}}def\ Three{\isaliteral{5F}{\isacharunderscore}}def\ Abs{\isaliteral{5F}{\isacharunderscore}}three{\isaliteral{5F}{\isacharunderscore}}inject\ three{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
 | 
|
1688  | 
\endisatagproof  | 
|
1689  | 
{\isafoldproof}%
 | 
|
1690  | 
%  | 
|
1691  | 
\isadelimproof  | 
|
1692  | 
\isanewline  | 
|
1693  | 
%  | 
|
1694  | 
\endisadelimproof  | 
|
1695  | 
\isanewline  | 
|
1696  | 
\isacommand{lemma}\isamarkupfalse%
 | 
|
1697  | 
\ three{\isaliteral{5F}{\isacharunderscore}}cases{\isaliteral{3A}{\isacharcolon}}\isanewline
 | 
|
1698  | 
\ \ \isakeyword{fixes}\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ three\ \isakeyword{obtains}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ One{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Two{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Three{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
 | 
|
1699  | 
%  | 
|
1700  | 
\isadelimproof  | 
|
1701  | 
\ \ %  | 
|
1702  | 
\endisadelimproof  | 
|
1703  | 
%  | 
|
1704  | 
\isatagproof  | 
|
1705  | 
\isacommand{by}\isamarkupfalse%
 | 
|
1706  | 
\ {\isaliteral{28}{\isacharparenleft}}cases\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}auto\ simp{\isaliteral{3A}{\isacharcolon}}\ One{\isaliteral{5F}{\isacharunderscore}}def\ Two{\isaliteral{5F}{\isacharunderscore}}def\ Three{\isaliteral{5F}{\isacharunderscore}}def\ Abs{\isaliteral{5F}{\isacharunderscore}}three{\isaliteral{5F}{\isacharunderscore}}inject\ three{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
 | 
|
1707  | 
\endisatagproof  | 
|
1708  | 
{\isafoldproof}%
 | 
|
1709  | 
%  | 
|
1710  | 
\isadelimproof  | 
|
1711  | 
%  | 
|
1712  | 
\endisadelimproof  | 
|
1713  | 
%  | 
|
1714  | 
\begin{isamarkuptext}%
 | 
|
1715  | 
Note that such trivial constructions are better done with  | 
|
1716  | 
  derived specification mechanisms such as \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}}:%
 | 
|
1717  | 
\end{isamarkuptext}%
 | 
|
1718  | 
\isamarkuptrue%  | 
|
1719  | 
\isacommand{datatype}\isamarkupfalse%
 | 
|
1720  | 
\ three{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ One{\isaliteral{27}{\isacharprime}}\ {\isaliteral{7C}{\isacharbar}}\ Two{\isaliteral{27}{\isacharprime}}\ {\isaliteral{7C}{\isacharbar}}\ Three{\isaliteral{27}{\isacharprime}}%
 | 
|
1721  | 
\begin{isamarkuptext}%
 | 
|
1722  | 
This avoids re-doing basic definitions and proofs from the  | 
|
1723  | 
  primitive \hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}} above.%
 | 
|
1724  | 
\end{isamarkuptext}%
 | 
|
1725  | 
\isamarkuptrue%  | 
|
1726  | 
%  | 
|
1727  | 
\isamarkupsection{Functorial structure of types%
 | 
|
1728  | 
}  | 
|
1729  | 
\isamarkuptrue%  | 
|
1730  | 
%  | 
|
1731  | 
\begin{isamarkuptext}%
 | 
|
1732  | 
\begin{matharray}{rcl}
 | 
|
1733  | 
    \indexdef{HOL}{command}{enriched\_type}\hypertarget{command.HOL.enriched-type}{\hyperlink{command.HOL.enriched-type}{\mbox{\isa{\isacommand{enriched{\isaliteral{5F}{\isacharunderscore}}type}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1734  | 
  \end{matharray}
 | 
|
| 26849 | 1735  | 
|
| 42908 | 1736  | 
  \begin{railoutput}
 | 
1737  | 
\rail@begin{2}{}
 | 
|
1738  | 
\rail@term{\hyperlink{command.HOL.enriched-type}{\mbox{\isa{\isacommand{enriched{\isaliteral{5F}{\isacharunderscore}}type}}}}}[]
 | 
|
1739  | 
\rail@bar  | 
|
1740  | 
\rail@nextbar{1}
 | 
|
1741  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
1742  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
|
1743  | 
\rail@endbar  | 
|
1744  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
|
1745  | 
\rail@end  | 
|
1746  | 
\end{railoutput}
 | 
|
1747  | 
||
1748  | 
||
1749  | 
  \begin{description}
 | 
|
| 26849 | 1750  | 
|
| 42908 | 1751  | 
  \item \hyperlink{command.HOL.enriched-type}{\mbox{\isa{\isacommand{enriched{\isaliteral{5F}{\isacharunderscore}}type}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}prefix{\isaliteral{3A}{\isacharcolon}}\ m{\isaliteral{22}{\isachardoublequote}}} allows to
 | 
1752  | 
prove and register properties about the functorial structure of type  | 
|
1753  | 
constructors. These properties then can be used by other packages  | 
|
1754  | 
to deal with those type constructors in certain type constructions.  | 
|
1755  | 
Characteristic theorems are noted in the current local theory. By  | 
|
1756  | 
default, they are prefixed with the base name of the type  | 
|
1757  | 
constructor, an explicit prefix can be given alternatively.  | 
|
1758  | 
||
1759  | 
  The given term \isa{{\isaliteral{22}{\isachardoublequote}}m{\isaliteral{22}{\isachardoublequote}}} is considered as \emph{mapper} for the
 | 
|
1760  | 
corresponding type constructor and must conform to the following  | 
|
1761  | 
type pattern:  | 
|
| 26849 | 1762  | 
|
| 42908 | 1763  | 
  \begin{matharray}{lll}
 | 
1764  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}m{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} &
 | 
|
1765  | 
      \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub k\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1766  | 
  \end{matharray}
 | 
|
1767  | 
||
1768  | 
  \noindent where \isa{t} is the type constructor, \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} are distinct
 | 
|
1769  | 
  type variables free in the local theory and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}},
 | 
|
1770  | 
  \ldots, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub k{\isaliteral{22}{\isachardoublequote}}} is a subsequence of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}, \ldots,
 | 
|
1771  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
1772  | 
||
1773  | 
  \end{description}%
 | 
|
| 26849 | 1774  | 
\end{isamarkuptext}%
 | 
1775  | 
\isamarkuptrue%  | 
|
1776  | 
%  | 
|
| 47821 | 1777  | 
\isamarkupsection{Transfer package%
 | 
1778  | 
}  | 
|
1779  | 
\isamarkuptrue%  | 
|
1780  | 
%  | 
|
1781  | 
\begin{isamarkuptext}%
 | 
|
1782  | 
\begin{matharray}{rcl}
 | 
|
1783  | 
    \indexdef{HOL}{method}{transfer}\hypertarget{method.HOL.transfer}{\hyperlink{method.HOL.transfer}{\mbox{\isa{transfer}}}} & : & \isa{method} \\
 | 
|
1784  | 
    \indexdef{HOL}{method}{transfer'}\hypertarget{method.HOL.transfer'}{\hyperlink{method.HOL.transfer'}{\mbox{\isa{transfer{\isaliteral{27}{\isacharprime}}}}}} & : & \isa{method} \\
 | 
|
1785  | 
    \indexdef{HOL}{method}{transfer\_prover}\hypertarget{method.HOL.transfer-prover}{\hyperlink{method.HOL.transfer-prover}{\mbox{\isa{transfer{\isaliteral{5F}{\isacharunderscore}}prover}}}} & : & \isa{method} \\
 | 
|
1786  | 
    \indexdef{HOL}{attribute}{transfer\_rule}\hypertarget{attribute.HOL.transfer-rule}{\hyperlink{attribute.HOL.transfer-rule}{\mbox{\isa{transfer{\isaliteral{5F}{\isacharunderscore}}rule}}}} & : & \isa{attribute} \\
 | 
|
1787  | 
    \indexdef{HOL}{attribute}{relator\_eq}\hypertarget{attribute.HOL.relator-eq}{\hyperlink{attribute.HOL.relator-eq}{\mbox{\isa{relator{\isaliteral{5F}{\isacharunderscore}}eq}}}} & : & \isa{attribute} \\
 | 
|
1788  | 
  \end{matharray}
 | 
|
1789  | 
||
1790  | 
  \begin{description}
 | 
|
1791  | 
||
1792  | 
  \item \hyperlink{method.HOL.transfer}{\mbox{\isa{transfer}}} method replaces the current subgoal
 | 
|
1793  | 
with a logically equivalent one that uses different types and  | 
|
1794  | 
constants. The replacement of types and constants is guided by the  | 
|
1795  | 
database of transfer rules. Goals are generalized over all free  | 
|
1796  | 
variables by default; this is necessary for variables whose types  | 
|
1797  | 
change, but can be overridden for specific variables with e.g.  | 
|
1798  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}transfer\ fixing{\isaliteral{3A}{\isacharcolon}}\ x\ y\ z{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
1799  | 
||
1800  | 
  \item \hyperlink{method.HOL.transfer'}{\mbox{\isa{transfer{\isaliteral{27}{\isacharprime}}}}} is a variant of \hyperlink{method.HOL.transfer}{\mbox{\isa{transfer}}} that allows replacing a subgoal with one that is
 | 
|
1801  | 
logically stronger (rather than equivalent). For example, a  | 
|
1802  | 
subgoal involving equality on a quotient type could be replaced  | 
|
1803  | 
with a subgoal involving equality (instead of the corresponding  | 
|
1804  | 
equivalence relation) on the underlying raw type.  | 
|
1805  | 
||
1806  | 
  \item \hyperlink{method.HOL.transfer-prover}{\mbox{\isa{transfer{\isaliteral{5F}{\isacharunderscore}}prover}}} method assists with proving
 | 
|
1807  | 
a transfer rule for a new constant, provided the constant is  | 
|
1808  | 
defined in terms of other constants that already have transfer  | 
|
1809  | 
rules. It should be applied after unfolding the constant  | 
|
1810  | 
definitions.  | 
|
1811  | 
||
1812  | 
  \item \hyperlink{attribute.HOL.transfer-rule}{\mbox{\isa{transfer{\isaliteral{5F}{\isacharunderscore}}rule}}} attribute maintains a
 | 
|
1813  | 
collection of transfer rules, which relate constants at two  | 
|
1814  | 
different types. Typical transfer rules may relate different type  | 
|
1815  | 
instances of the same polymorphic constant, or they may relate an  | 
|
1816  | 
operation on a raw type to a corresponding operation on an  | 
|
1817  | 
abstract type (quotient or subtype). For example:  | 
|
1818  | 
||
1819  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ list{\isaliteral{5F}{\isacharunderscore}}all{\isadigit{2}}\ A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ list{\isaliteral{5F}{\isacharunderscore}}all{\isadigit{2}}\ B{\isaliteral{29}{\isacharparenright}}\ map\ map{\isaliteral{22}{\isachardoublequote}}}\\
 | 
|
1820  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}cr{\isaliteral{5F}{\isacharunderscore}}int\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ cr{\isaliteral{5F}{\isacharunderscore}}int\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ cr{\isaliteral{5F}{\isacharunderscore}}int{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{2C}{\isacharcomma}}y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}u{\isaliteral{2C}{\isacharcomma}}v{\isaliteral{29}{\isacharparenright}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{2B}{\isacharplus}}u{\isaliteral{2C}{\isacharcomma}}\ y{\isaliteral{2B}{\isacharplus}}v{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ plus{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1821  | 
||
1822  | 
Lemmas involving predicates on relations can also be registered  | 
|
1823  | 
using the same attribute. For example:  | 
|
1824  | 
||
1825  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}bi{\isaliteral{5F}{\isacharunderscore}}unique\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}list{\isaliteral{5F}{\isacharunderscore}}all{\isadigit{2}}\ A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ op\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}\ distinct\ distinct{\isaliteral{22}{\isachardoublequote}}}\\
 | 
|
1826  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}bi{\isaliteral{5F}{\isacharunderscore}}unique\ A{\isaliteral{3B}{\isacharsemicolon}}\ bi{\isaliteral{5F}{\isacharunderscore}}unique\ B{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ bi{\isaliteral{5F}{\isacharunderscore}}unique\ {\isaliteral{28}{\isacharparenleft}}prod{\isaliteral{5F}{\isacharunderscore}}rel\ A\ B{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1827  | 
||
1828  | 
  \item \hyperlink{attribute.HOL.relator-eq}{\mbox{\isa{relator{\isaliteral{5F}{\isacharunderscore}}eq}}} attribute collects identity laws
 | 
|
1829  | 
    for relators of various type constructors, e.g. \isa{{\isaliteral{22}{\isachardoublequote}}list{\isaliteral{5F}{\isacharunderscore}}all{\isadigit{2}}\ {\isaliteral{28}{\isacharparenleft}}op\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}op\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. The \hyperlink{method.HOL.transfer}{\mbox{\isa{transfer}}} method uses these
 | 
|
1830  | 
lemmas to infer transfer rules for non-polymorphic constants on  | 
|
1831  | 
the fly.  | 
|
1832  | 
||
1833  | 
  \end{description}%
 | 
|
1834  | 
\end{isamarkuptext}%
 | 
|
1835  | 
\isamarkuptrue%  | 
|
1836  | 
%  | 
|
| 47802 | 1837  | 
\isamarkupsection{Lifting package%
 | 
1838  | 
}  | 
|
1839  | 
\isamarkuptrue%  | 
|
1840  | 
%  | 
|
1841  | 
\begin{isamarkuptext}%
 | 
|
1842  | 
\begin{matharray}{rcl}
 | 
|
1843  | 
    \indexdef{HOL}{command}{setup\_lifting}\hypertarget{command.HOL.setup-lifting}{\hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}}\\
 | 
|
1844  | 
    \indexdef{HOL}{command}{lift\_definition}\hypertarget{command.HOL.lift-definition}{\hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}\\
 | 
|
1845  | 
    \indexdef{HOL}{command}{print\_quotmaps}\hypertarget{command.HOL.print-quotmaps}{\hyperlink{command.HOL.print-quotmaps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmaps}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\
 | 
|
1846  | 
    \indexdef{HOL}{command}{print\_quotients}\hypertarget{command.HOL.print-quotients}{\hyperlink{command.HOL.print-quotients}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotients}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\
 | 
|
1847  | 
    \indexdef{HOL}{attribute}{quot\_map}\hypertarget{attribute.HOL.quot-map}{\hyperlink{attribute.HOL.quot-map}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}map}}}} & : & \isa{attribute} \\
 | 
|
1848  | 
    \indexdef{HOL}{attribute}{invariant\_commute}\hypertarget{attribute.HOL.invariant-commute}{\hyperlink{attribute.HOL.invariant-commute}{\mbox{\isa{invariant{\isaliteral{5F}{\isacharunderscore}}commute}}}} & : & \isa{attribute} \\
 | 
|
1849  | 
  \end{matharray}
 | 
|
1850  | 
||
1851  | 
  \begin{railoutput}
 | 
|
1852  | 
\rail@begin{5}{}
 | 
|
1853  | 
\rail@term{\hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}}}[]
 | 
|
1854  | 
\rail@bar  | 
|
1855  | 
\rail@nextbar{1}
 | 
|
1856  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
|
1857  | 
\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}abs{\isaliteral{5F}{\isacharunderscore}}code}}[]
 | 
|
1858  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
|
1859  | 
\rail@endbar  | 
|
1860  | 
\rail@cr{3}
 | 
|
1861  | 
\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[]
 | 
|
1862  | 
\rail@bar  | 
|
1863  | 
\rail@nextbar{4}
 | 
|
1864  | 
\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[]
 | 
|
1865  | 
\rail@endbar  | 
|
1866  | 
\rail@end  | 
|
1867  | 
\end{railoutput}
 | 
|
1868  | 
||
1869  | 
||
1870  | 
  \begin{railoutput}
 | 
|
1871  | 
\rail@begin{4}{}
 | 
|
1872  | 
\rail@term{\hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}}}[]
 | 
|
1873  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
1874  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
|
1875  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
|
1876  | 
\rail@bar  | 
|
1877  | 
\rail@nextbar{1}
 | 
|
1878  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
|
1879  | 
\rail@endbar  | 
|
1880  | 
\rail@cr{3}
 | 
|
1881  | 
\rail@term{\isa{is}}[]
 | 
|
1882  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
|
1883  | 
\rail@end  | 
|
1884  | 
\end{railoutput}
 | 
|
1885  | 
||
1886  | 
||
1887  | 
  \begin{description}
 | 
|
1888  | 
||
| 47859 | 1889  | 
  \item \hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}} Sets up the Lifting package
 | 
1890  | 
to work with a user-defined type. The user must provide either a  | 
|
1891  | 
    quotient theorem \isa{{\isaliteral{22}{\isachardoublequote}}Quotient\ R\ Abs\ Rep\ T{\isaliteral{22}{\isachardoublequote}}} or a
 | 
|
1892  | 
    type_definition theorem \isa{{\isaliteral{22}{\isachardoublequote}}type{\isaliteral{5F}{\isacharunderscore}}definition\ Rep\ Abs\ A{\isaliteral{22}{\isachardoublequote}}}.  The
 | 
|
1893  | 
package configures transfer rules for equality and quantifiers on  | 
|
1894  | 
    the type, and sets up the \indexdef{HOL}{command}{lift\_definition}\hypertarget{command.HOL.lift-definition}{\hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}}}
 | 
|
1895  | 
command to work with the type. In the case of a quotient theorem,  | 
|
1896  | 
    an optional theorem \isa{{\isaliteral{22}{\isachardoublequote}}reflp\ R{\isaliteral{22}{\isachardoublequote}}} can be provided as a second
 | 
|
1897  | 
argument. This allows the package to generate stronger transfer  | 
|
1898  | 
rules.  | 
|
1899  | 
||
1900  | 
    \hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}} is called automatically if a
 | 
|
1901  | 
    quotient type is defined by the command \hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}} from the Quotient package.
 | 
|
1902  | 
||
1903  | 
    If \hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}} is called with a
 | 
|
1904  | 
type_definition theorem, the abstract type implicitly defined by  | 
|
1905  | 
the theorem is declared as an abstract type in the code  | 
|
1906  | 
    generator. This allows \hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}} to
 | 
|
1907  | 
register (generated) code certificate theorems as abstract code  | 
|
1908  | 
    equations in the code generator.  The option \isa{{\isaliteral{22}{\isachardoublequote}}no{\isaliteral{5F}{\isacharunderscore}}abs{\isaliteral{5F}{\isacharunderscore}}code{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1909  | 
    of the command \hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}} can turn off that
 | 
|
1910  | 
behavior and causes that code certificate theorems generated by  | 
|
1911  | 
    \hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}} are not registred as abstract
 | 
|
1912  | 
code equations.  | 
|
1913  | 
||
1914  | 
  \item \hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}} \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\ is\ t{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1915  | 
    Defines a new function \isa{f} with an abstract type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}}
 | 
|
1916  | 
    in terms of a corresponding operation \isa{t} on a
 | 
|
1917  | 
    representation type.  The term \isa{t} doesn't have to be
 | 
|
1918  | 
necessarily a constant but it can be any term.  | 
|
| 47802 | 1919  | 
|
1920  | 
Users must discharge a respectfulness proof obligation when each  | 
|
| 47859 | 1921  | 
    constant is defined. For a type copy, i.e. a typedef with \isa{UNIV}, the proof is discharged automatically. The obligation is
 | 
| 47802 | 1922  | 
presented in a user-friendly, readable form. A respectfulness  | 
| 47859 | 1923  | 
    theorem in the standard format \isa{f{\isaliteral{2E}{\isachardot}}rsp} and a transfer rule
 | 
1924  | 
    \isa{f{\isaliteral{2E}{\isachardot}}tranfer} for the Transfer package are generated by the
 | 
|
1925  | 
package.  | 
|
| 47802 | 1926  | 
|
1927  | 
Integration with code_abstype: For typedefs (e.g. subtypes  | 
|
| 47859 | 1928  | 
    corresponding to a datatype invariant, such as dlist), \hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}} generates a code certificate theorem
 | 
1929  | 
    \isa{f{\isaliteral{2E}{\isachardot}}rep{\isaliteral{5F}{\isacharunderscore}}eq} and sets up code generation for each constant.
 | 
|
| 47802 | 1930  | 
|
1931  | 
  \item \hyperlink{command.HOL.print-quotmaps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmaps}}}} prints stored quotient map
 | 
|
| 47859 | 1932  | 
theorems.  | 
1933  | 
||
1934  | 
  \item \hyperlink{command.HOL.print-quotients}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotients}}}} prints stored quotient
 | 
|
1935  | 
theorems.  | 
|
1936  | 
||
1937  | 
  \item \hyperlink{attribute.HOL.quot-map}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}map}}} registers a quotient map
 | 
|
1938  | 
theorem. For examples see \verb|~~/src/HOL/Library/Quotient_List.thy| or other Quotient_*.thy  | 
|
1939  | 
files.  | 
|
1940  | 
||
1941  | 
  \item \hyperlink{attribute.HOL.invariant-commute}{\mbox{\isa{invariant{\isaliteral{5F}{\isacharunderscore}}commute}}} registers a theorem which
 | 
|
1942  | 
    shows a relationship between the constant \isa{Lifting{\isaliteral{2E}{\isachardot}}invariant} (used for internal encoding of proper subtypes)
 | 
|
1943  | 
    and a relator.  Such theorems allows the package to hide \isa{Lifting{\isaliteral{2E}{\isachardot}}invariant} from a user in a user-readable form of a
 | 
|
1944  | 
respectfulness theorem. For examples see \verb|~~/src/HOL/Library/Quotient_List.thy| or other Quotient_*.thy  | 
|
1945  | 
files.  | 
|
| 47802 | 1946  | 
|
1947  | 
  \end{description}%
 | 
|
1948  | 
\end{isamarkuptext}%
 | 
|
1949  | 
\isamarkuptrue%  | 
|
1950  | 
%  | 
|
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1951  | 
\isamarkupsection{Quotient types%
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1952  | 
}  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1953  | 
\isamarkuptrue%  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1954  | 
%  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1955  | 
\begin{isamarkuptext}%
 | 
| 46280 | 1956  | 
\begin{matharray}{rcl}
 | 
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1957  | 
    \indexdef{HOL}{command}{quotient\_type}\hypertarget{command.HOL.quotient-type}{\hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}\\
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1958  | 
    \indexdef{HOL}{command}{quotient\_definition}\hypertarget{command.HOL.quotient-definition}{\hyperlink{command.HOL.quotient-definition}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}definition}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}\\
 | 
| 47802 | 1959  | 
    \indexdef{HOL}{command}{print\_quotmapsQ3}\hypertarget{command.HOL.print-quotmapsQ3}{\hyperlink{command.HOL.print-quotmapsQ3}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmapsQ{\isadigit{3}}}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\
 | 
1960  | 
    \indexdef{HOL}{command}{print\_quotientsQ3}\hypertarget{command.HOL.print-quotientsQ3}{\hyperlink{command.HOL.print-quotientsQ3}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotientsQ{\isadigit{3}}}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\
 | 
|
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1961  | 
    \indexdef{HOL}{command}{print\_quotconsts}\hypertarget{command.HOL.print-quotconsts}{\hyperlink{command.HOL.print-quotconsts}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotconsts}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\
 | 
| 
46447
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1962  | 
    \indexdef{HOL}{method}{lifting}\hypertarget{method.HOL.lifting}{\hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}}} & : & \isa{method} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1963  | 
    \indexdef{HOL}{method}{lifting\_setup}\hypertarget{method.HOL.lifting-setup}{\hyperlink{method.HOL.lifting-setup}{\mbox{\isa{lifting{\isaliteral{5F}{\isacharunderscore}}setup}}}} & : & \isa{method} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1964  | 
    \indexdef{HOL}{method}{descending}\hypertarget{method.HOL.descending}{\hyperlink{method.HOL.descending}{\mbox{\isa{descending}}}} & : & \isa{method} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1965  | 
    \indexdef{HOL}{method}{descending\_setup}\hypertarget{method.HOL.descending-setup}{\hyperlink{method.HOL.descending-setup}{\mbox{\isa{descending{\isaliteral{5F}{\isacharunderscore}}setup}}}} & : & \isa{method} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1966  | 
    \indexdef{HOL}{method}{partiality\_descending}\hypertarget{method.HOL.partiality-descending}{\hyperlink{method.HOL.partiality-descending}{\mbox{\isa{partiality{\isaliteral{5F}{\isacharunderscore}}descending}}}} & : & \isa{method} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1967  | 
    \indexdef{HOL}{method}{partiality\_descending\_setup}\hypertarget{method.HOL.partiality-descending-setup}{\hyperlink{method.HOL.partiality-descending-setup}{\mbox{\isa{partiality{\isaliteral{5F}{\isacharunderscore}}descending{\isaliteral{5F}{\isacharunderscore}}setup}}}} & : & \isa{method} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1968  | 
    \indexdef{HOL}{method}{regularize}\hypertarget{method.HOL.regularize}{\hyperlink{method.HOL.regularize}{\mbox{\isa{regularize}}}} & : & \isa{method} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1969  | 
    \indexdef{HOL}{method}{injection}\hypertarget{method.HOL.injection}{\hyperlink{method.HOL.injection}{\mbox{\isa{injection}}}} & : & \isa{method} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1970  | 
    \indexdef{HOL}{method}{cleaning}\hypertarget{method.HOL.cleaning}{\hyperlink{method.HOL.cleaning}{\mbox{\isa{cleaning}}}} & : & \isa{method} \\
 | 
| 
46448
 
f1201fac7398
more specification of the quotient package in IsarRef
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46447 
diff
changeset
 | 
1971  | 
    \indexdef{HOL}{attribute}{quot\_thm}\hypertarget{attribute.HOL.quot-thm}{\hyperlink{attribute.HOL.quot-thm}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}thm}}}} & : & \isa{attribute} \\
 | 
| 
46447
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1972  | 
    \indexdef{HOL}{attribute}{quot\_lifted}\hypertarget{attribute.HOL.quot-lifted}{\hyperlink{attribute.HOL.quot-lifted}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}lifted}}}} & : & \isa{attribute} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1973  | 
    \indexdef{HOL}{attribute}{quot\_respect}\hypertarget{attribute.HOL.quot-respect}{\hyperlink{attribute.HOL.quot-respect}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}respect}}}} & : & \isa{attribute} \\
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
1974  | 
    \indexdef{HOL}{attribute}{quot\_preserve}\hypertarget{attribute.HOL.quot-preserve}{\hyperlink{attribute.HOL.quot-preserve}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}preserve}}}} & : & \isa{attribute} \\
 | 
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1975  | 
  \end{matharray}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1976  | 
|
| 46280 | 1977  | 
The quotient package defines a new quotient type given a raw type  | 
1978  | 
and a partial equivalence relation. It also includes automation for  | 
|
1979  | 
transporting definitions and theorems. It can automatically produce  | 
|
1980  | 
definitions and theorems on the quotient type, given the  | 
|
1981  | 
corresponding constants and facts on the raw type.  | 
|
1982  | 
||
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1983  | 
  \begin{railoutput}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1984  | 
\rail@begin{2}{}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1985  | 
\rail@term{\hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1986  | 
\rail@plus  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1987  | 
\rail@nont{\isa{spec}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1988  | 
\rail@nextplus{1}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1989  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1990  | 
\rail@endplus  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1991  | 
\rail@end  | 
| 45678 | 1992  | 
\rail@begin{8}{\isa{spec}}
 | 
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1993  | 
\rail@nont{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1994  | 
\rail@bar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1995  | 
\rail@nextbar{1}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1996  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1997  | 
\rail@endbar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1998  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1999  | 
\rail@cr{3}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2000  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2001  | 
\rail@term{\isa{{\isaliteral{2F}{\isacharslash}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2002  | 
\rail@bar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2003  | 
\rail@nextbar{4}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2004  | 
\rail@term{\isa{partial}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2005  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2006  | 
\rail@endbar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2007  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 45678 | 2008  | 
\rail@cr{6}
 | 
2009  | 
\rail@bar  | 
|
2010  | 
\rail@nextbar{7}
 | 
|
2011  | 
\rail@term{\isa{\isakeyword{morphisms}}}[]
 | 
|
2012  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
2013  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
|
2014  | 
\rail@endbar  | 
|
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2015  | 
\rail@end  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2016  | 
\end{railoutput}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2017  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2018  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2019  | 
  \begin{railoutput}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2020  | 
\rail@begin{4}{}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2021  | 
\rail@term{\hyperlink{command.HOL.quotient-definition}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}definition}}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2022  | 
\rail@bar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2023  | 
\rail@nextbar{1}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2024  | 
\rail@nont{\isa{constdecl}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2025  | 
\rail@endbar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2026  | 
\rail@bar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2027  | 
\rail@nextbar{1}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2028  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2029  | 
\rail@endbar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2030  | 
\rail@cr{3}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2031  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2032  | 
\rail@term{\isa{is}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2033  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2034  | 
\rail@end  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2035  | 
\rail@begin{2}{\isa{constdecl}}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2036  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2037  | 
\rail@bar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2038  | 
\rail@nextbar{1}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2039  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2040  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2041  | 
\rail@endbar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2042  | 
\rail@bar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2043  | 
\rail@nextbar{1}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2044  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2045  | 
\rail@endbar  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2046  | 
\rail@end  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2047  | 
\end{railoutput}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2048  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2049  | 
|
| 
46447
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2050  | 
  \begin{railoutput}
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2051  | 
\rail@begin{2}{}
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2052  | 
\rail@term{\hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}}}[]
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2053  | 
\rail@bar  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2054  | 
\rail@nextbar{1}
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2055  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2056  | 
\rail@endbar  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2057  | 
\rail@end  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2058  | 
\rail@begin{2}{}
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2059  | 
\rail@term{\hyperlink{method.HOL.lifting-setup}{\mbox{\isa{lifting{\isaliteral{5F}{\isacharunderscore}}setup}}}}[]
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2060  | 
\rail@bar  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2061  | 
\rail@nextbar{1}
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2062  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2063  | 
\rail@endbar  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2064  | 
\rail@end  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2065  | 
\end{railoutput}
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2066  | 
|
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2067  | 
|
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2068  | 
  \begin{description}
 | 
| 
46447
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2069  | 
|
| 46457 | 2070  | 
  \item \hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}} defines quotient types. The
 | 
2071  | 
  injection from a quotient type to a raw type is called \isa{rep{\isaliteral{5F}{\isacharunderscore}}t}, its inverse \isa{abs{\isaliteral{5F}{\isacharunderscore}}t} unless explicit \hyperlink{keyword.HOL.morphisms}{\mbox{\isa{\isakeyword{morphisms}}}} specification provides alternative names. \hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}} requires the user to prove that the relation
 | 
|
2072  | 
  is an equivalence relation (predicate \isa{equivp}), unless the
 | 
|
2073  | 
  user specifies explicitely \isa{partial} in which case the
 | 
|
2074  | 
  obligation is \isa{part{\isaliteral{5F}{\isacharunderscore}}equivp}.  A quotient defined with \isa{partial} is weaker in the sense that less things can be proved
 | 
|
| 
46447
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2075  | 
automatically.  | 
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2076  | 
|
| 46280 | 2077  | 
  \item \hyperlink{command.HOL.quotient-definition}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}definition}}}} defines a constant on
 | 
2078  | 
the quotient type.  | 
|
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2079  | 
|
| 47802 | 2080  | 
  \item \hyperlink{command.HOL.print-quotmapsQ3}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmapsQ{\isadigit{3}}}}}} prints quotient map
 | 
| 46280 | 2081  | 
functions.  | 
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2082  | 
|
| 47802 | 2083  | 
  \item \hyperlink{command.HOL.print-quotientsQ3}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotientsQ{\isadigit{3}}}}}} prints quotients.
 | 
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2084  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2085  | 
  \item \hyperlink{command.HOL.print-quotconsts}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotconsts}}}} prints quotient constants.
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2086  | 
|
| 
46447
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2087  | 
  \item \hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}} and \hyperlink{method.HOL.lifting-setup}{\mbox{\isa{lifting{\isaliteral{5F}{\isacharunderscore}}setup}}}
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2088  | 
methods match the current goal with the given raw theorem to be  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2089  | 
lifted producing three new subgoals: regularization, injection and  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2090  | 
    cleaning subgoals. \hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}} tries to apply the
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2091  | 
heuristics for automatically solving these three subgoals and  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2092  | 
leaves only the subgoals unsolved by the heuristics to the user as  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2093  | 
    opposed to \hyperlink{method.HOL.lifting-setup}{\mbox{\isa{lifting{\isaliteral{5F}{\isacharunderscore}}setup}}} which leaves the three
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2094  | 
subgoals unsolved.  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2095  | 
|
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2096  | 
  \item \hyperlink{method.HOL.descending}{\mbox{\isa{descending}}} and \hyperlink{method.HOL.descending-setup}{\mbox{\isa{descending{\isaliteral{5F}{\isacharunderscore}}setup}}} try to guess a raw statement that would lift
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2097  | 
to the current subgoal. Such statement is assumed as a new subgoal  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2098  | 
    and \hyperlink{method.HOL.descending}{\mbox{\isa{descending}}} continues in the same way as
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2099  | 
    \hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}} does. \hyperlink{method.HOL.descending}{\mbox{\isa{descending}}} tries
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2100  | 
to solve the arising regularization, injection and cleaning  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2101  | 
    subgoals with the analoguous method \hyperlink{method.HOL.descending-setup}{\mbox{\isa{descending{\isaliteral{5F}{\isacharunderscore}}setup}}} which leaves the four unsolved subgoals.
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2102  | 
|
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2103  | 
  \item \hyperlink{method.HOL.partiality-descending}{\mbox{\isa{partiality{\isaliteral{5F}{\isacharunderscore}}descending}}} finds the regularized
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2104  | 
theorem that would lift to the current subgoal, lifts it and  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2105  | 
leaves as a subgoal. This method can be used with partial  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2106  | 
equivalence quotients where the non regularized statements would  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2107  | 
    not be true. \hyperlink{method.HOL.partiality-descending-setup}{\mbox{\isa{partiality{\isaliteral{5F}{\isacharunderscore}}descending{\isaliteral{5F}{\isacharunderscore}}setup}}} leaves
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2108  | 
the injection and cleaning subgoals unchanged.  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2109  | 
|
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2110  | 
  \item \hyperlink{method.HOL.regularize}{\mbox{\isa{regularize}}} applies the regularization
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2111  | 
heuristics to the current subgoal.  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2112  | 
|
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2113  | 
  \item \hyperlink{method.HOL.injection}{\mbox{\isa{injection}}} applies the injection heuristics
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2114  | 
to the current goal using the stored quotient respectfulness  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2115  | 
theorems.  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2116  | 
|
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2117  | 
  \item \hyperlink{method.HOL.cleaning}{\mbox{\isa{cleaning}}} applies the injection cleaning
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2118  | 
heuristics to the current subgoal using the stored quotient  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2119  | 
preservation theorems.  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2120  | 
|
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2121  | 
  \item \hyperlink{attribute.HOL.quot-lifted}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}lifted}}} attribute tries to
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2122  | 
automatically transport the theorem to the quotient type.  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2123  | 
The attribute uses all the defined quotients types and quotient  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2124  | 
constants often producing undesired results or theorems that  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2125  | 
cannot be lifted.  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2126  | 
|
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2127  | 
  \item \hyperlink{attribute.HOL.quot-respect}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}respect}}} and \hyperlink{attribute.HOL.quot-preserve}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}preserve}}} attributes declare a theorem as a respectfulness
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2128  | 
and preservation theorem respectively. These are stored in the  | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2129  | 
    local theory store and used by the \hyperlink{method.HOL.injection}{\mbox{\isa{injection}}}
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2130  | 
    and \hyperlink{method.HOL.cleaning}{\mbox{\isa{cleaning}}} methods respectively.
 | 
| 
 
f37da60a8cc6
specification of the quotient package
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46035 
diff
changeset
 | 
2131  | 
|
| 
46448
 
f1201fac7398
more specification of the quotient package in IsarRef
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46447 
diff
changeset
 | 
2132  | 
  \item \hyperlink{attribute.HOL.quot-thm}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}thm}}} declares that a certain theorem
 | 
| 
 
f1201fac7398
more specification of the quotient package in IsarRef
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46447 
diff
changeset
 | 
2133  | 
is a quotient extension theorem. Quotient extension theorems  | 
| 
 
f1201fac7398
more specification of the quotient package in IsarRef
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46447 
diff
changeset
 | 
2134  | 
allow for quotienting inside container types. Given a polymorphic  | 
| 
 
f1201fac7398
more specification of the quotient package in IsarRef
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46447 
diff
changeset
 | 
2135  | 
type that serves as a container, a map function defined for this  | 
| 
47349
 
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
 
bulwahn 
parents: 
46705 
diff
changeset
 | 
2136  | 
    container using \hyperlink{command.HOL.enriched-type}{\mbox{\isa{\isacommand{enriched{\isaliteral{5F}{\isacharunderscore}}type}}}} and a relation
 | 
| 
46448
 
f1201fac7398
more specification of the quotient package in IsarRef
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46447 
diff
changeset
 | 
2137  | 
map defined for for the container type, the quotient extension  | 
| 
47349
 
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
 
bulwahn 
parents: 
46705 
diff
changeset
 | 
2138  | 
    theorem should be \isa{{\isaliteral{22}{\isachardoublequote}}Quotient{\isadigit{3}}\ R\ Abs\ Rep\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Quotient{\isadigit{3}}\ {\isaliteral{28}{\isacharparenleft}}rel{\isaliteral{5F}{\isacharunderscore}}map\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}map\ Abs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}map\ Rep{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. Quotient extension theorems
 | 
| 
46448
 
f1201fac7398
more specification of the quotient package in IsarRef
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46447 
diff
changeset
 | 
2139  | 
are stored in a database and are used all the steps of lifting  | 
| 
 
f1201fac7398
more specification of the quotient package in IsarRef
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46447 
diff
changeset
 | 
2140  | 
theorems.  | 
| 
 
f1201fac7398
more specification of the quotient package in IsarRef
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents: 
46447 
diff
changeset
 | 
2141  | 
|
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2142  | 
  \end{description}%
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2143  | 
\end{isamarkuptext}%
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2144  | 
\isamarkuptrue%  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
2145  | 
%  | 
| 43994 | 2146  | 
\isamarkupsection{Coercive subtyping%
 | 
2147  | 
}  | 
|
2148  | 
\isamarkuptrue%  | 
|
2149  | 
%  | 
|
2150  | 
\begin{isamarkuptext}%
 | 
|
2151  | 
\begin{matharray}{rcl}
 | 
|
2152  | 
    \indexdef{HOL}{attribute}{coercion}\hypertarget{attribute.HOL.coercion}{\hyperlink{attribute.HOL.coercion}{\mbox{\isa{coercion}}}} & : & \isa{attribute} \\
 | 
|
2153  | 
    \indexdef{HOL}{attribute}{coercion\_enabled}\hypertarget{attribute.HOL.coercion-enabled}{\hyperlink{attribute.HOL.coercion-enabled}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}enabled}}}} & : & \isa{attribute} \\
 | 
|
2154  | 
    \indexdef{HOL}{attribute}{coercion\_map}\hypertarget{attribute.HOL.coercion-map}{\hyperlink{attribute.HOL.coercion-map}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}map}}}} & : & \isa{attribute} \\
 | 
|
2155  | 
  \end{matharray}
 | 
|
2156  | 
||
| 46280 | 2157  | 
Coercive subtyping allows the user to omit explicit type  | 
2158  | 
  conversions, also called \emph{coercions}.  Type inference will add
 | 
|
2159  | 
them as necessary when parsing a term. See  | 
|
2160  | 
  \cite{traytel-berghofer-nipkow-2011} for details.
 | 
|
2161  | 
||
| 43994 | 2162  | 
  \begin{railoutput}
 | 
2163  | 
\rail@begin{2}{}
 | 
|
2164  | 
\rail@term{\hyperlink{attribute.HOL.coercion}{\mbox{\isa{coercion}}}}[]
 | 
|
2165  | 
\rail@bar  | 
|
2166  | 
\rail@nextbar{1}
 | 
|
2167  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
|
2168  | 
\rail@endbar  | 
|
2169  | 
\rail@end  | 
|
2170  | 
\end{railoutput}
 | 
|
2171  | 
||
2172  | 
  \begin{railoutput}
 | 
|
2173  | 
\rail@begin{2}{}
 | 
|
2174  | 
\rail@term{\hyperlink{attribute.HOL.coercion-map}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}map}}}}[]
 | 
|
2175  | 
\rail@bar  | 
|
2176  | 
\rail@nextbar{1}
 | 
|
2177  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
|
2178  | 
\rail@endbar  | 
|
2179  | 
\rail@end  | 
|
2180  | 
\end{railoutput}
 | 
|
2181  | 
||
2182  | 
||
2183  | 
  \begin{description}
 | 
|
2184  | 
||
2185  | 
  \item \hyperlink{attribute.HOL.coercion}{\mbox{\isa{coercion}}}~\isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} registers a new
 | 
|
| 46280 | 2186  | 
  coercion function \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} where \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} and
 | 
2187  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} are type constructors without arguments.  Coercions are
 | 
|
2188  | 
composed by the inference algorithm if needed. Note that the type  | 
|
2189  | 
inference algorithm is complete only if the registered coercions  | 
|
2190  | 
form a lattice.  | 
|
| 43994 | 2191  | 
|
| 46280 | 2192  | 
  \item \hyperlink{attribute.HOL.coercion-map}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}map}}}~\isa{{\isaliteral{22}{\isachardoublequote}}map{\isaliteral{22}{\isachardoublequote}}} registers a
 | 
2193  | 
new map function to lift coercions through type constructors. The  | 
|
2194  | 
  function \isa{{\isaliteral{22}{\isachardoublequote}}map{\isaliteral{22}{\isachardoublequote}}} must conform to the following type pattern
 | 
|
| 43994 | 2195  | 
|
2196  | 
  \begin{matharray}{lll}
 | 
|
2197  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}map{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} &
 | 
|
2198  | 
      \isa{{\isaliteral{22}{\isachardoublequote}}f\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ f\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2199  | 
  \end{matharray}
 | 
|
2200  | 
||
| 46280 | 2201  | 
  where \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}} is a type constructor and \isa{{\isaliteral{22}{\isachardoublequote}}f\isaliteral{5C3C5E697375623E}{}\isactrlisub i{\isaliteral{22}{\isachardoublequote}}} is of type
 | 
2202  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub i{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub i{\isaliteral{22}{\isachardoublequote}}}.  Registering a map function
 | 
|
2203  | 
overwrites any existing map function for this particular type  | 
|
2204  | 
constructor.  | 
|
| 43994 | 2205  | 
|
2206  | 
  \item \hyperlink{attribute.HOL.coercion-enabled}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}enabled}}} enables the coercion
 | 
|
2207  | 
inference algorithm.  | 
|
2208  | 
||
2209  | 
  \end{description}%
 | 
|
2210  | 
\end{isamarkuptext}%
 | 
|
2211  | 
\isamarkuptrue%  | 
|
2212  | 
%  | 
|
| 26849 | 2213  | 
\isamarkupsection{Arithmetic proof support%
 | 
2214  | 
}  | 
|
2215  | 
\isamarkuptrue%  | 
|
2216  | 
%  | 
|
2217  | 
\begin{isamarkuptext}%
 | 
|
2218  | 
\begin{matharray}{rcl}
 | 
|
| 28788 | 2219  | 
    \indexdef{HOL}{method}{arith}\hypertarget{method.HOL.arith}{\hyperlink{method.HOL.arith}{\mbox{\isa{arith}}}} & : & \isa{method} \\
 | 
| 30863 | 2220  | 
    \indexdef{HOL}{attribute}{arith}\hypertarget{attribute.HOL.arith}{\hyperlink{attribute.HOL.arith}{\mbox{\isa{arith}}}} & : & \isa{attribute} \\
 | 
| 40406 | 2221  | 
    \indexdef{HOL}{attribute}{arith\_split}\hypertarget{attribute.HOL.arith-split}{\hyperlink{attribute.HOL.arith-split}{\mbox{\isa{arith{\isaliteral{5F}{\isacharunderscore}}split}}}} & : & \isa{attribute} \\
 | 
| 26849 | 2222  | 
  \end{matharray}
 | 
2223  | 
||
| 46280 | 2224  | 
  \begin{description}
 | 
| 26849 | 2225  | 
|
| 46280 | 2226  | 
  \item \hyperlink{method.HOL.arith}{\mbox{\isa{arith}}} decides linear arithmetic problems (on
 | 
2227  | 
  types \isa{nat}, \isa{int}, \isa{real}).  Any current facts
 | 
|
2228  | 
are inserted into the goal before running the procedure.  | 
|
| 26849 | 2229  | 
|
| 46280 | 2230  | 
  \item \hyperlink{attribute.HOL.arith}{\mbox{\isa{arith}}} declares facts that are supplied to
 | 
2231  | 
the arithmetic provers implicitly.  | 
|
2232  | 
||
2233  | 
  \item \hyperlink{attribute.HOL.arith-split}{\mbox{\isa{arith{\isaliteral{5F}{\isacharunderscore}}split}}} attribute declares case split
 | 
|
| 30865 | 2234  | 
  rules to be expanded before \hyperlink{method.HOL.arith}{\mbox{\isa{arith}}} is invoked.
 | 
| 30863 | 2235  | 
|
| 46280 | 2236  | 
  \end{description}
 | 
2237  | 
||
2238  | 
Note that a simpler (but faster) arithmetic prover is already  | 
|
2239  | 
invoked by the Simplifier.%  | 
|
| 26849 | 2240  | 
\end{isamarkuptext}%
 | 
2241  | 
\isamarkuptrue%  | 
|
2242  | 
%  | 
|
| 30172 | 2243  | 
\isamarkupsection{Intuitionistic proof search%
 | 
2244  | 
}  | 
|
2245  | 
\isamarkuptrue%  | 
|
2246  | 
%  | 
|
2247  | 
\begin{isamarkuptext}%
 | 
|
2248  | 
\begin{matharray}{rcl}
 | 
|
2249  | 
    \indexdef{HOL}{method}{iprover}\hypertarget{method.HOL.iprover}{\hyperlink{method.HOL.iprover}{\mbox{\isa{iprover}}}} & : & \isa{method} \\
 | 
|
2250  | 
  \end{matharray}
 | 
|
2251  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2252  | 
  \begin{railoutput}
 | 
| 42662 | 2253  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2254  | 
\rail@term{\hyperlink{method.HOL.iprover}{\mbox{\isa{iprover}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2255  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2256  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2257  | 
\rail@cnont{\hyperlink{syntax.rulemod}{\mbox{\isa{rulemod}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2258  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2259  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2260  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2261  | 
|
| 30172 | 2262  | 
|
| 46280 | 2263  | 
  \begin{description}
 | 
2264  | 
||
2265  | 
  \item \hyperlink{method.HOL.iprover}{\mbox{\isa{iprover}}} performs intuitionistic proof search,
 | 
|
2266  | 
depending on specifically declared rules from the context, or given  | 
|
2267  | 
as explicit arguments. Chained facts are inserted into the goal  | 
|
2268  | 
before commencing proof search.  | 
|
| 35613 | 2269  | 
|
| 30172 | 2270  | 
  Rules need to be classified as \hyperlink{attribute.Pure.intro}{\mbox{\isa{intro}}},
 | 
2271  | 
  \hyperlink{attribute.Pure.elim}{\mbox{\isa{elim}}}, or \hyperlink{attribute.Pure.dest}{\mbox{\isa{dest}}}; here the
 | 
|
| 40406 | 2272  | 
  ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}}'' indicator refers to ``safe'' rules, which may be
 | 
| 30172 | 2273  | 
applied aggressively (without considering back-tracking later).  | 
| 40406 | 2274  | 
  Rules declared with ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' are ignored in proof search (the
 | 
| 42626 | 2275  | 
  single-step \hyperlink{method.Pure.rule}{\mbox{\isa{rule}}} method still observes these).  An
 | 
| 30172 | 2276  | 
explicit weight annotation may be given as well; otherwise the  | 
| 46280 | 2277  | 
number of rule premises will be taken into account here.  | 
2278  | 
||
2279  | 
  \end{description}%
 | 
|
| 30172 | 2280  | 
\end{isamarkuptext}%
 | 
2281  | 
\isamarkuptrue%  | 
|
2282  | 
%  | 
|
| 
43578
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2283  | 
\isamarkupsection{Model Elimination and Resolution%
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2284  | 
}  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2285  | 
\isamarkuptrue%  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2286  | 
%  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2287  | 
\begin{isamarkuptext}%
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2288  | 
\begin{matharray}{rcl}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2289  | 
    \indexdef{HOL}{method}{meson}\hypertarget{method.HOL.meson}{\hyperlink{method.HOL.meson}{\mbox{\isa{meson}}}} & : & \isa{method} \\
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2290  | 
    \indexdef{HOL}{method}{metis}\hypertarget{method.HOL.metis}{\hyperlink{method.HOL.metis}{\mbox{\isa{metis}}}} & : & \isa{method} \\
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2291  | 
  \end{matharray}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2292  | 
|
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2293  | 
  \begin{railoutput}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2294  | 
\rail@begin{2}{}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2295  | 
\rail@term{\hyperlink{method.HOL.meson}{\mbox{\isa{meson}}}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2296  | 
\rail@bar  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2297  | 
\rail@nextbar{1}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2298  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2299  | 
\rail@endbar  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2300  | 
\rail@end  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2301  | 
\rail@begin{5}{}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2302  | 
\rail@term{\hyperlink{method.HOL.metis}{\mbox{\isa{metis}}}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2303  | 
\rail@bar  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2304  | 
\rail@nextbar{1}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2305  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2306  | 
\rail@bar  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2307  | 
\rail@term{\isa{partial{\isaliteral{5F}{\isacharunderscore}}types}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2308  | 
\rail@nextbar{2}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2309  | 
\rail@term{\isa{full{\isaliteral{5F}{\isacharunderscore}}types}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2310  | 
\rail@nextbar{3}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2311  | 
\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}types}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2312  | 
\rail@nextbar{4}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2313  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2314  | 
\rail@endbar  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2315  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2316  | 
\rail@endbar  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2317  | 
\rail@bar  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2318  | 
\rail@nextbar{1}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2319  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2320  | 
\rail@endbar  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2321  | 
\rail@end  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2322  | 
\end{railoutput}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2323  | 
|
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2324  | 
|
| 46280 | 2325  | 
  \begin{description}
 | 
2326  | 
||
2327  | 
  \item \hyperlink{method.HOL.meson}{\mbox{\isa{meson}}} implements Loveland's model elimination
 | 
|
2328  | 
  procedure \cite{loveland-78}.  See \verb|~~/src/HOL/ex/Meson_Test.thy| for examples.
 | 
|
| 
43578
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2329  | 
|
| 46280 | 2330  | 
  \item \hyperlink{method.HOL.metis}{\mbox{\isa{metis}}} combines ordered resolution and ordered
 | 
2331  | 
paramodulation to find first-order (or mildly higher-order) proofs.  | 
|
2332  | 
The first optional argument specifies a type encoding; see the  | 
|
2333  | 
  Sledgehammer manual \cite{isabelle-sledgehammer} for details.  The
 | 
|
2334  | 
directory \verb|~~/src/HOL/Metis_Examples| contains several small  | 
|
2335  | 
  theories developed to a large extent using \hyperlink{method.HOL.metis}{\mbox{\isa{metis}}}.
 | 
|
2336  | 
||
2337  | 
  \end{description}%
 | 
|
| 
43578
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2338  | 
\end{isamarkuptext}%
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2339  | 
\isamarkuptrue%  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
2340  | 
%  | 
| 30172 | 2341  | 
\isamarkupsection{Coherent Logic%
 | 
2342  | 
}  | 
|
2343  | 
\isamarkuptrue%  | 
|
2344  | 
%  | 
|
2345  | 
\begin{isamarkuptext}%
 | 
|
2346  | 
\begin{matharray}{rcl}
 | 
|
2347  | 
    \indexdef{HOL}{method}{coherent}\hypertarget{method.HOL.coherent}{\hyperlink{method.HOL.coherent}{\mbox{\isa{coherent}}}} & : & \isa{method} \\
 | 
|
2348  | 
  \end{matharray}
 | 
|
2349  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2350  | 
  \begin{railoutput}
 | 
| 42662 | 2351  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2352  | 
\rail@term{\hyperlink{method.HOL.coherent}{\mbox{\isa{coherent}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2353  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2354  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2355  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2356  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2357  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2358  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2359  | 
|
| 30172 | 2360  | 
|
| 46280 | 2361  | 
  \begin{description}
 | 
2362  | 
||
2363  | 
  \item \hyperlink{method.HOL.coherent}{\mbox{\isa{coherent}}} solves problems of \emph{Coherent
 | 
|
2364  | 
  Logic} \cite{Bezem-Coquand:2005}, which covers applications in
 | 
|
2365  | 
confluence theory, lattice theory and projective geometry. See  | 
|
2366  | 
\verb|~~/src/HOL/ex/Coherent.thy| for some examples.  | 
|
2367  | 
||
2368  | 
  \end{description}%
 | 
|
| 30172 | 2369  | 
\end{isamarkuptext}%
 | 
2370  | 
\isamarkuptrue%  | 
|
2371  | 
%  | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2372  | 
\isamarkupsection{Proving propositions%
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2373  | 
}  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2374  | 
\isamarkuptrue%  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2375  | 
%  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2376  | 
\begin{isamarkuptext}%
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2377  | 
In addition to the standard proof methods, a number of diagnosis  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2378  | 
tools search for proofs and provide an Isar proof snippet on success.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2379  | 
These tools are available via the following commands.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2380  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2381  | 
  \begin{matharray}{rcl}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2382  | 
    \indexdef{HOL}{command}{solve\_direct}\hypertarget{command.HOL.solve-direct}{\hyperlink{command.HOL.solve-direct}{\mbox{\isa{\isacommand{solve{\isaliteral{5F}{\isacharunderscore}}direct}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2383  | 
    \indexdef{HOL}{command}{try}\hypertarget{command.HOL.try}{\hyperlink{command.HOL.try}{\mbox{\isa{\isacommand{try}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 46705 | 2384  | 
    \indexdef{HOL}{command}{try0}\hypertarget{command.HOL.try0}{\hyperlink{command.HOL.try0}{\mbox{\isa{\isacommand{try{\isadigit{0}}}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2385  | 
    \indexdef{HOL}{command}{sledgehammer}\hypertarget{command.HOL.sledgehammer}{\hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2386  | 
    \indexdef{HOL}{command}{sledgehammer\_params}\hypertarget{command.HOL.sledgehammer-params}{\hyperlink{command.HOL.sledgehammer-params}{\mbox{\isa{\isacommand{sledgehammer{\isaliteral{5F}{\isacharunderscore}}params}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2387  | 
  \end{matharray}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2388  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2389  | 
  \begin{railoutput}
 | 
| 43040 | 2390  | 
\rail@begin{1}{}
 | 
2391  | 
\rail@term{\hyperlink{command.HOL.try}{\mbox{\isa{\isacommand{try}}}}}[]
 | 
|
2392  | 
\rail@end  | 
|
| 42662 | 2393  | 
\rail@begin{6}{}
 | 
| 46705 | 2394  | 
\rail@term{\hyperlink{command.HOL.try0}{\mbox{\isa{\isacommand{try{\isadigit{0}}}}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2395  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2396  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2397  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2398  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2399  | 
\rail@term{\isa{simp}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2400  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2401  | 
\rail@term{\isa{intro}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2402  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2403  | 
\rail@term{\isa{elim}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2404  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2405  | 
\rail@term{\isa{dest}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2406  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2407  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2408  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2409  | 
\rail@nextplus{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2410  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2411  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2412  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2413  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2414  | 
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2415  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2416  | 
\rail@end  | 
| 42662 | 2417  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2418  | 
\rail@term{\hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2419  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2420  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2421  | 
\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2422  | 
\rail@nont{\isa{args}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2423  | 
\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2424  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2425  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2426  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2427  | 
\rail@nont{\isa{facts}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2428  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2429  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2430  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2431  | 
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2432  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2433  | 
\rail@end  | 
| 42662 | 2434  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2435  | 
\rail@term{\hyperlink{command.HOL.sledgehammer-params}{\mbox{\isa{\isacommand{sledgehammer{\isaliteral{5F}{\isacharunderscore}}params}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2436  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2437  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2438  | 
\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2439  | 
\rail@nont{\isa{args}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2440  | 
\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2441  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2442  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2443  | 
\rail@begin{2}{\isa{args}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2444  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2445  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2446  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2447  | 
\rail@nont{\isa{value}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2448  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2449  | 
\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2450  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2451  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2452  | 
\rail@begin{5}{\isa{facts}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2453  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2454  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2455  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2456  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2457  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2458  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2459  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2460  | 
\rail@term{\isa{add}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2461  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2462  | 
\rail@term{\isa{del}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2463  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2464  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2465  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2466  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2467  | 
\rail@nextplus{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2468  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2469  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2470  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2471  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2472  | 
\end{railoutput}
 | 
| 43040 | 2473  | 
% FIXME check args "value"  | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2474  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2475  | 
  \begin{description}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2476  | 
|
| 46283 | 2477  | 
  \item \hyperlink{command.HOL.solve-direct}{\mbox{\isa{\isacommand{solve{\isaliteral{5F}{\isacharunderscore}}direct}}}} checks whether the current
 | 
2478  | 
subgoals can be solved directly by an existing theorem. Duplicate  | 
|
2479  | 
lemmas can be detected in this way.  | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2480  | 
|
| 46705 | 2481  | 
  \item \hyperlink{command.HOL.try0}{\mbox{\isa{\isacommand{try{\isadigit{0}}}}}} attempts to prove a subgoal
 | 
| 46283 | 2482  | 
  using a combination of standard proof methods (\hyperlink{method.auto}{\mbox{\isa{auto}}},
 | 
2483  | 
  \hyperlink{method.simp}{\mbox{\isa{simp}}}, \hyperlink{method.blast}{\mbox{\isa{blast}}}, etc.).  Additional facts supplied
 | 
|
2484  | 
  via \isa{{\isaliteral{22}{\isachardoublequote}}simp{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}intro{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}elim{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}, and \isa{{\isaliteral{22}{\isachardoublequote}}dest{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} are passed to the appropriate proof methods.
 | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2485  | 
|
| 
43914
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
2486  | 
  \item \hyperlink{command.HOL.try}{\mbox{\isa{\isacommand{try}}}} attempts to prove or disprove a subgoal
 | 
| 46705 | 2487  | 
  using a combination of provers and disprovers (\hyperlink{command.HOL.solve-direct}{\mbox{\isa{\isacommand{solve{\isaliteral{5F}{\isacharunderscore}}direct}}}}, \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}, \hyperlink{command.HOL.try0}{\mbox{\isa{\isacommand{try{\isadigit{0}}}}}}, \hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}}, \hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}}).
 | 
| 
43914
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
2488  | 
|
| 46283 | 2489  | 
  \item \hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}} attempts to prove a subgoal
 | 
2490  | 
using external automatic provers (resolution provers and SMT  | 
|
2491  | 
  solvers). See the Sledgehammer manual \cite{isabelle-sledgehammer}
 | 
|
2492  | 
for details.  | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2493  | 
|
| 46283 | 2494  | 
  \item \hyperlink{command.HOL.sledgehammer-params}{\mbox{\isa{\isacommand{sledgehammer{\isaliteral{5F}{\isacharunderscore}}params}}}} changes \hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}} configuration options persistently.
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2495  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2496  | 
  \end{description}%
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2497  | 
\end{isamarkuptext}%
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2498  | 
\isamarkuptrue%  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2499  | 
%  | 
| 31913 | 2500  | 
\isamarkupsection{Checking and refuting propositions%
 | 
2501  | 
}  | 
|
2502  | 
\isamarkuptrue%  | 
|
2503  | 
%  | 
|
2504  | 
\begin{isamarkuptext}%
 | 
|
2505  | 
Identifying incorrect propositions usually involves evaluation of  | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2506  | 
particular assignments and systematic counterexample search. This  | 
| 31913 | 2507  | 
is supported by the following commands.  | 
2508  | 
||
2509  | 
  \begin{matharray}{rcl}
 | 
|
| 40406 | 2510  | 
    \indexdef{HOL}{command}{value}\hypertarget{command.HOL.value}{\hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
45409
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2511  | 
    \indexdef{HOL}{command}{values}\hypertarget{command.HOL.values}{\hyperlink{command.HOL.values}{\mbox{\isa{\isacommand{values}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 40406 | 2512  | 
    \indexdef{HOL}{command}{quickcheck}\hypertarget{command.HOL.quickcheck}{\hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2513  | 
    \indexdef{HOL}{command}{refute}\hypertarget{command.HOL.refute}{\hyperlink{command.HOL.refute}{\mbox{\isa{\isacommand{refute}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2514  | 
    \indexdef{HOL}{command}{nitpick}\hypertarget{command.HOL.nitpick}{\hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2515  | 
    \indexdef{HOL}{command}{quickcheck\_params}\hypertarget{command.HOL.quickcheck-params}{\hyperlink{command.HOL.quickcheck-params}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}params}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2516  | 
    \indexdef{HOL}{command}{refute\_params}\hypertarget{command.HOL.refute-params}{\hyperlink{command.HOL.refute-params}{\mbox{\isa{\isacommand{refute{\isaliteral{5F}{\isacharunderscore}}params}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
45943
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2517  | 
    \indexdef{HOL}{command}{nitpick\_params}\hypertarget{command.HOL.nitpick-params}{\hyperlink{command.HOL.nitpick-params}{\mbox{\isa{\isacommand{nitpick{\isaliteral{5F}{\isacharunderscore}}params}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
46592
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2518  | 
    \indexdef{HOL}{command}{quickcheck\_generator}\hypertarget{command.HOL.quickcheck-generator}{\hyperlink{command.HOL.quickcheck-generator}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}generator}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2519  | 
    \indexdef{HOL}{command}{find\_unused\_assms}\hypertarget{command.HOL.find-unused-assms}{\hyperlink{command.HOL.find-unused-assms}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}unused{\isaliteral{5F}{\isacharunderscore}}assms}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}
 | 
| 31913 | 2520  | 
  \end{matharray}
 | 
2521  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2522  | 
  \begin{railoutput}
 | 
| 42662 | 2523  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2524  | 
\rail@term{\hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2525  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2526  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2527  | 
\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[]
 | 
| 46628 | 2528  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2529  | 
\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2530  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2531  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2532  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2533  | 
\rail@nont{\isa{modes}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2534  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2535  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2536  | 
\rail@end  | 
| 
45409
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2537  | 
\rail@begin{2}{}
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2538  | 
\rail@term{\hyperlink{command.HOL.values}{\mbox{\isa{\isacommand{values}}}}}[]
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2539  | 
\rail@bar  | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2540  | 
\rail@nextbar{1}
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2541  | 
\rail@nont{\isa{modes}}[]
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2542  | 
\rail@endbar  | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2543  | 
\rail@bar  | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2544  | 
\rail@nextbar{1}
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2545  | 
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2546  | 
\rail@endbar  | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2547  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2548  | 
\rail@end  | 
| 42662 | 2549  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2550  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2551  | 
\rail@term{\hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2552  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2553  | 
\rail@term{\hyperlink{command.HOL.refute}{\mbox{\isa{\isacommand{refute}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2554  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2555  | 
\rail@term{\hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2556  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2557  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2558  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2559  | 
\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2560  | 
\rail@nont{\isa{args}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2561  | 
\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2562  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2563  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2564  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2565  | 
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2566  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2567  | 
\rail@end  | 
| 42662 | 2568  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2569  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2570  | 
\rail@term{\hyperlink{command.HOL.quickcheck-params}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}params}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2571  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2572  | 
\rail@term{\hyperlink{command.HOL.refute-params}{\mbox{\isa{\isacommand{refute{\isaliteral{5F}{\isacharunderscore}}params}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2573  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2574  | 
\rail@term{\hyperlink{command.HOL.nitpick-params}{\mbox{\isa{\isacommand{nitpick{\isaliteral{5F}{\isacharunderscore}}params}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2575  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2576  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2577  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2578  | 
\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2579  | 
\rail@nont{\isa{args}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2580  | 
\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2581  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2582  | 
\rail@end  | 
| 
45943
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2583  | 
\rail@begin{4}{}
 | 
| 
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2584  | 
\rail@term{\hyperlink{command.HOL.quickcheck-generator}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}generator}}}}}[]
 | 
| 46628 | 2585  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
45943
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2586  | 
\rail@cr{2}
 | 
| 
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2587  | 
\rail@term{\isa{operations{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2588  | 
\rail@plus  | 
| 
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2589  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2590  | 
\rail@nextplus{3}
 | 
| 
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2591  | 
\rail@endplus  | 
| 
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2592  | 
\rail@end  | 
| 
46592
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2593  | 
\rail@begin{2}{}
 | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2594  | 
\rail@term{\hyperlink{command.HOL.find-unused-assms}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}unused{\isaliteral{5F}{\isacharunderscore}}assms}}}}}[]
 | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2595  | 
\rail@bar  | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2596  | 
\rail@nextbar{1}
 | 
| 46628 | 2597  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
46592
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2598  | 
\rail@endbar  | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2599  | 
\rail@end  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2600  | 
\rail@begin{2}{\isa{modes}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2601  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2602  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2603  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2604  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2605  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2606  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2607  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2608  | 
\rail@begin{2}{\isa{args}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2609  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2610  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2611  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2612  | 
\rail@nont{\isa{value}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2613  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2614  | 
\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2615  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2616  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2617  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2618  | 
% FIXME check "value"  | 
| 31913 | 2619  | 
|
2620  | 
  \begin{description}
 | 
|
2621  | 
||
2622  | 
  \item \hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}~\isa{t} evaluates and prints a
 | 
|
| 46283 | 2623  | 
  term; optionally \isa{modes} can be specified, which are appended
 | 
2624  | 
  to the current print mode; see \secref{sec:print-modes}.
 | 
|
2625  | 
Internally, the evaluation is performed by registered evaluators,  | 
|
2626  | 
which are invoked sequentially until a result is returned.  | 
|
2627  | 
Alternatively a specific evaluator can be selected using square  | 
|
2628  | 
brackets; typical evaluators use the current set of code equations  | 
|
2629  | 
  to normalize and include \isa{simp} for fully symbolic evaluation
 | 
|
2630  | 
  using the simplifier, \isa{nbe} for \emph{normalization by
 | 
|
2631  | 
  evaluation} and \emph{code} for code generation in SML.
 | 
|
| 31913 | 2632  | 
|
| 46283 | 2633  | 
  \item \hyperlink{command.HOL.values}{\mbox{\isa{\isacommand{values}}}}~\isa{t} enumerates a set
 | 
2634  | 
comprehension by evaluation and prints its values up to the given  | 
|
2635  | 
  number of solutions; optionally \isa{modes} can be specified,
 | 
|
2636  | 
which are appended to the current print mode; see  | 
|
2637  | 
  \secref{sec:print-modes}.
 | 
|
| 
45409
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
2638  | 
|
| 31913 | 2639  | 
  \item \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}} tests the current goal for
 | 
| 46283 | 2640  | 
counterexamples using a series of assignments for its free  | 
2641  | 
variables; by default the first subgoal is tested, an other can be  | 
|
2642  | 
selected explicitly using an optional goal index. Assignments can  | 
|
2643  | 
be chosen exhausting the search space upto a given size, or using a  | 
|
2644  | 
fixed number of random assignments in the search space, or exploring  | 
|
2645  | 
the search space symbolically using narrowing. By default,  | 
|
2646  | 
quickcheck uses exhaustive testing. A number of configuration  | 
|
2647  | 
  options are supported for \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}, notably:
 | 
|
| 31913 | 2648  | 
|
2649  | 
    \begin{description}
 | 
|
2650  | 
||
| 
43914
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
2651  | 
    \item[\isa{tester}] specifies which testing approach to apply.
 | 
| 46283 | 2652  | 
    There are three testers, \isa{exhaustive}, \isa{random}, and
 | 
2653  | 
    \isa{narrowing}.  An unknown configuration option is treated as
 | 
|
2654  | 
    an argument to tester, making \isa{{\isaliteral{22}{\isachardoublequote}}tester\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}} optional.  When
 | 
|
2655  | 
multiple testers are given, these are applied in parallel. If no  | 
|
2656  | 
tester is specified, quickcheck uses the testers that are set  | 
|
2657  | 
    active, i.e., configurations \hyperlink{attribute.quickcheck-exhaustive-active}{\mbox{\isa{quickcheck{\isaliteral{5F}{\isacharunderscore}}exhaustive{\isaliteral{5F}{\isacharunderscore}}active}}}, \hyperlink{attribute.quickcheck-random-active}{\mbox{\isa{quickcheck{\isaliteral{5F}{\isacharunderscore}}random{\isaliteral{5F}{\isacharunderscore}}active}}}, \hyperlink{attribute.quickcheck-narrowing-active}{\mbox{\isa{quickcheck{\isaliteral{5F}{\isacharunderscore}}narrowing{\isaliteral{5F}{\isacharunderscore}}active}}} are set to true.
 | 
|
2658  | 
||
| 40254 | 2659  | 
    \item[\isa{size}] specifies the maximum size of the search space
 | 
2660  | 
for assignment values.  | 
|
2661  | 
||
| 
45758
 
6210c350d88b
documenting the genuine_only option in quickcheck;
 
bulwahn 
parents: 
45701 
diff
changeset
 | 
2662  | 
    \item[\isa{genuine{\isaliteral{5F}{\isacharunderscore}}only}] sets quickcheck only to return genuine
 | 
| 46283 | 2663  | 
counterexample, but not potentially spurious counterexamples due  | 
2664  | 
to underspecified functions.  | 
|
| 
46498
 
2754784e9153
adding documentation for abort_potential option in quickcheck
 
bulwahn 
parents: 
46494 
diff
changeset
 | 
2665  | 
|
| 
 
2754784e9153
adding documentation for abort_potential option in quickcheck
 
bulwahn 
parents: 
46494 
diff
changeset
 | 
2666  | 
    \item[\isa{abort{\isaliteral{5F}{\isacharunderscore}}potential}] sets quickcheck to abort once it
 | 
| 
 
2754784e9153
adding documentation for abort_potential option in quickcheck
 
bulwahn 
parents: 
46494 
diff
changeset
 | 
2667  | 
found a potentially spurious counterexample and to not continue  | 
| 
 
2754784e9153
adding documentation for abort_potential option in quickcheck
 
bulwahn 
parents: 
46494 
diff
changeset
 | 
2668  | 
to search for a further genuine counterexample.  | 
| 
 
2754784e9153
adding documentation for abort_potential option in quickcheck
 
bulwahn 
parents: 
46494 
diff
changeset
 | 
2669  | 
    For this option to be effective, the \isa{genuine{\isaliteral{5F}{\isacharunderscore}}only} option
 | 
| 
 
2754784e9153
adding documentation for abort_potential option in quickcheck
 
bulwahn 
parents: 
46494 
diff
changeset
 | 
2670  | 
must be set to false.  | 
| 47859 | 2671  | 
|
| 42123 | 2672  | 
    \item[\isa{eval}] takes a term or a list of terms and evaluates
 | 
| 46283 | 2673  | 
these terms under the variable assignment found by quickcheck.  | 
| 48159 | 2674  | 
This option is currently only supported by the default  | 
2675  | 
(exhaustive) tester.  | 
|
| 42123 | 2676  | 
|
| 40254 | 2677  | 
    \item[\isa{iterations}] sets how many sets of assignments are
 | 
2678  | 
generated for each particular size.  | 
|
2679  | 
||
| 40406 | 2680  | 
    \item[\isa{no{\isaliteral{5F}{\isacharunderscore}}assms}] specifies whether assumptions in
 | 
| 40254 | 2681  | 
structured proofs should be ignored.  | 
2682  | 
||
| 
47349
 
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
 
bulwahn 
parents: 
46705 
diff
changeset
 | 
2683  | 
    \item[\isa{locale}] specifies how to process conjectures in
 | 
| 
 
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
 
bulwahn 
parents: 
46705 
diff
changeset
 | 
2684  | 
a locale context, i.e., they can be interpreted or expanded.  | 
| 
 
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
 
bulwahn 
parents: 
46705 
diff
changeset
 | 
2685  | 
The option is a whitespace-separated list of the two words  | 
| 
 
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
 
bulwahn 
parents: 
46705 
diff
changeset
 | 
2686  | 
    \isa{interpret} and \isa{expand}. The list determines the
 | 
| 47859 | 2687  | 
order they are employed. The default setting is to first use  | 
| 
47349
 
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
 
bulwahn 
parents: 
46705 
diff
changeset
 | 
2688  | 
interpretations and then test the expanded conjecture.  | 
| 
 
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
 
bulwahn 
parents: 
46705 
diff
changeset
 | 
2689  | 
The option is only provided as attribute declaration, but not  | 
| 47859 | 2690  | 
as parameter to the command.  | 
| 
47349
 
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
 
bulwahn 
parents: 
46705 
diff
changeset
 | 
2691  | 
|
| 40254 | 2692  | 
    \item[\isa{timeout}] sets the time limit in seconds.
 | 
| 31913 | 2693  | 
|
| 40406 | 2694  | 
    \item[\isa{default{\isaliteral{5F}{\isacharunderscore}}type}] sets the type(s) generally used to
 | 
| 40254 | 2695  | 
instantiate type variables.  | 
2696  | 
||
2697  | 
    \item[\isa{report}] if set quickcheck reports how many tests
 | 
|
2698  | 
fulfilled the preconditions.  | 
|
| 31913 | 2699  | 
|
| 
46592
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2700  | 
    \item[\isa{use{\isaliteral{5F}{\isacharunderscore}}subtype}] if set quickcheck automatically lifts
 | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2701  | 
conjectures to registered subtypes if possible, and tests the  | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2702  | 
lifted conjecture.  | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2703  | 
|
| 
45766
 
46046d8e9659
updating documentation about quiet and verbose options in quickcheck
 
bulwahn 
parents: 
45758 
diff
changeset
 | 
2704  | 
    \item[\isa{quiet}] if set quickcheck does not output anything
 | 
| 
 
46046d8e9659
updating documentation about quiet and verbose options in quickcheck
 
bulwahn 
parents: 
45758 
diff
changeset
 | 
2705  | 
while testing.  | 
| 47859 | 2706  | 
|
| 46283 | 2707  | 
    \item[\isa{verbose}] if set quickcheck informs about the current
 | 
2708  | 
size and cardinality while testing.  | 
|
| 40254 | 2709  | 
|
2710  | 
    \item[\isa{expect}] can be used to check if the user's
 | 
|
| 40406 | 2711  | 
    expectation was met (\isa{no{\isaliteral{5F}{\isacharunderscore}}expectation}, \isa{no{\isaliteral{5F}{\isacharunderscore}}counterexample}, or \isa{counterexample}).
 | 
| 
35351
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
34172 
diff
changeset
 | 
2712  | 
|
| 31913 | 2713  | 
    \end{description}
 | 
2714  | 
||
| 46283 | 2715  | 
These option can be given within square brackets.  | 
| 31913 | 2716  | 
|
| 46283 | 2717  | 
  \item \hyperlink{command.HOL.quickcheck-params}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}params}}}} changes \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}} configuration options persistently.
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2718  | 
|
| 
45943
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2719  | 
  \item \hyperlink{command.HOL.quickcheck-generator}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}generator}}}} creates random and
 | 
| 46283 | 2720  | 
exhaustive value generators for a given type and operations. It  | 
2721  | 
generates values by using the operations as if they were  | 
|
2722  | 
constructors of that type.  | 
|
| 
45943
 
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
 
bulwahn 
parents: 
45839 
diff
changeset
 | 
2723  | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2724  | 
  \item \hyperlink{command.HOL.refute}{\mbox{\isa{\isacommand{refute}}}} tests the current goal for
 | 
| 46283 | 2725  | 
counterexamples using a reduction to SAT. The following  | 
2726  | 
configuration options are supported:  | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2727  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2728  | 
    \begin{description}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2729  | 
|
| 46283 | 2730  | 
    \item[\isa{minsize}] specifies the minimum size (cardinality) of
 | 
2731  | 
the models to search for.  | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2732  | 
|
| 46283 | 2733  | 
    \item[\isa{maxsize}] specifies the maximum size (cardinality) of
 | 
2734  | 
    the models to search for. Nonpositive values mean \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E66696E6974793E}{\isasyminfinity}}{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2735  | 
|
| 46283 | 2736  | 
    \item[\isa{maxvars}] specifies the maximum number of Boolean
 | 
2737  | 
variables to use when transforming the term into a propositional  | 
|
2738  | 
    formula.  Nonpositive values mean \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E66696E6974793E}{\isasyminfinity}}{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2739  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2740  | 
    \item[\isa{satsolver}] specifies the SAT solver to use.
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2741  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2742  | 
    \item[\isa{no{\isaliteral{5F}{\isacharunderscore}}assms}] specifies whether assumptions in
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2743  | 
structured proofs should be ignored.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2744  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2745  | 
    \item[\isa{maxtime}] sets the time limit in seconds.
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2746  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2747  | 
    \item[\isa{expect}] can be used to check if the user's
 | 
| 46283 | 2748  | 
    expectation was met (\isa{genuine}, \isa{potential}, \isa{none}, or \isa{unknown}).
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2749  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2750  | 
    \end{description}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2751  | 
|
| 46283 | 2752  | 
These option can be given within square brackets.  | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2753  | 
|
| 46283 | 2754  | 
  \item \hyperlink{command.HOL.refute-params}{\mbox{\isa{\isacommand{refute{\isaliteral{5F}{\isacharunderscore}}params}}}} changes \hyperlink{command.HOL.refute}{\mbox{\isa{\isacommand{refute}}}} configuration options persistently.
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2755  | 
|
| 46283 | 2756  | 
  \item \hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}} tests the current goal for
 | 
2757  | 
counterexamples using a reduction to first-order relational  | 
|
2758  | 
  logic. See the Nitpick manual \cite{isabelle-nitpick} for details.
 | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
2759  | 
|
| 46283 | 2760  | 
  \item \hyperlink{command.HOL.nitpick-params}{\mbox{\isa{\isacommand{nitpick{\isaliteral{5F}{\isacharunderscore}}params}}}} changes \hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}} configuration options persistently.
 | 
| 31913 | 2761  | 
|
| 
46592
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2762  | 
  \item \hyperlink{command.HOL.find-unused-assms}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}unused{\isaliteral{5F}{\isacharunderscore}}assms}}}} finds potentially superfluous
 | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2763  | 
assumptions in theorems using quickcheck.  | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2764  | 
It takes the theory name to be checked for superfluous assumptions as  | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2765  | 
optional argument. If not provided, it checks the current theory.  | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2766  | 
Options to the internal quickcheck invocations can be changed with  | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2767  | 
common configuration declarations.  | 
| 
 
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
 
bulwahn 
parents: 
46525 
diff
changeset
 | 
2768  | 
|
| 31913 | 2769  | 
  \end{description}%
 | 
2770  | 
\end{isamarkuptext}%
 | 
|
2771  | 
\isamarkuptrue%  | 
|
2772  | 
%  | 
|
| 28788 | 2773  | 
\isamarkupsection{Unstructured case analysis and induction \label{sec:hol-induct-tac}%
 | 
| 26849 | 2774  | 
}  | 
2775  | 
\isamarkuptrue%  | 
|
2776  | 
%  | 
|
2777  | 
\begin{isamarkuptext}%
 | 
|
| 27124 | 2778  | 
The following tools of Isabelle/HOL support cases analysis and  | 
2779  | 
induction in unstructured tactic scripts; see also  | 
|
2780  | 
  \secref{sec:cases-induct} for proper Isar versions of similar ideas.
 | 
|
| 26849 | 2781  | 
|
2782  | 
  \begin{matharray}{rcl}
 | 
|
| 40406 | 2783  | 
    \indexdef{HOL}{method}{case\_tac}\hypertarget{method.HOL.case-tac}{\hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\
 | 
2784  | 
    \indexdef{HOL}{method}{induct\_tac}\hypertarget{method.HOL.induct-tac}{\hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\
 | 
|
2785  | 
    \indexdef{HOL}{method}{ind\_cases}\hypertarget{method.HOL.ind-cases}{\hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\
 | 
|
2786  | 
    \indexdef{HOL}{command}{inductive\_cases}\hypertarget{command.HOL.inductive-cases}{\hyperlink{command.HOL.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 26849 | 2787  | 
  \end{matharray}
 | 
2788  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2789  | 
  \begin{railoutput}
 | 
| 42662 | 2790  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2791  | 
\rail@term{\hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2792  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2793  | 
\rail@nextbar{1}
 | 
| 42705 | 2794  | 
\rail@nont{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2795  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2796  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2797  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2798  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2799  | 
\rail@nont{\isa{rule}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2800  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2801  | 
\rail@end  | 
| 42662 | 2802  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2803  | 
\rail@term{\hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2804  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2805  | 
\rail@nextbar{1}
 | 
| 42705 | 2806  | 
\rail@nont{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2807  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2808  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2809  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2810  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2811  | 
\rail@nont{\hyperlink{syntax.insts}{\mbox{\isa{insts}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2812  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2813  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2814  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2815  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2816  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2817  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2818  | 
\rail@nont{\isa{rule}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2819  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2820  | 
\rail@end  | 
| 42662 | 2821  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2822  | 
\rail@term{\hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2823  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2824  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2825  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2826  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2827  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2828  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2829  | 
\rail@term{\isa{\isakeyword{for}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2830  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2831  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2832  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2833  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2834  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2835  | 
\rail@end  | 
| 42662 | 2836  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2837  | 
\rail@term{\hyperlink{command.HOL.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2838  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2839  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2840  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2841  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2842  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2843  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2844  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2845  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2846  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2847  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2848  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2849  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2850  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2851  | 
\rail@begin{1}{\isa{rule}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2852  | 
\rail@term{\isa{rule}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2853  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2854  | 
\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2855  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2856  | 
\end{railoutput}
 | 
| 26849 | 2857  | 
|
2858  | 
||
| 28788 | 2859  | 
  \begin{description}
 | 
| 26849 | 2860  | 
|
| 40406 | 2861  | 
  \item \hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}} and \hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}} admit
 | 
| 28788 | 2862  | 
to reason about inductive types. Rules are selected according to  | 
2863  | 
  the declarations by the \hyperlink{attribute.cases}{\mbox{\isa{cases}}} and \hyperlink{attribute.induct}{\mbox{\isa{induct}}}
 | 
|
2864  | 
  attributes, cf.\ \secref{sec:cases-induct}.  The \hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}} package already takes care of this.
 | 
|
| 27124 | 2865  | 
|
2866  | 
These unstructured tactics feature both goal addressing and dynamic  | 
|
| 26849 | 2867  | 
  instantiation.  Note that named rule cases are \emph{not} provided
 | 
| 27124 | 2868  | 
  as would be by the proper \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} proof
 | 
| 40406 | 2869  | 
  methods (see \secref{sec:cases-induct}).  Unlike the \hyperlink{method.induct}{\mbox{\isa{induct}}} method, \hyperlink{method.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}} does not handle structured rule
 | 
| 27124 | 2870  | 
statements, only the compact object-logic conclusion of the subgoal  | 
2871  | 
being addressed.  | 
|
| 42123 | 2872  | 
|
| 40406 | 2873  | 
  \item \hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}} and \hyperlink{command.HOL.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}} provide an interface to the internal \verb|mk_cases| operation.  Rules are simplified in an unrestricted
 | 
| 26861 | 2874  | 
forward manner.  | 
| 26849 | 2875  | 
|
| 40406 | 2876  | 
  While \hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}} is a proof method to apply the
 | 
2877  | 
  result immediately as elimination rules, \hyperlink{command.HOL.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}} provides case split theorems at the theory level
 | 
|
2878  | 
  for later use.  The \hyperlink{keyword.for}{\mbox{\isa{\isakeyword{for}}}} argument of the \hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}} method allows to specify a list of variables that should
 | 
|
| 26849 | 2879  | 
be generalized before applying the resulting rule.  | 
2880  | 
||
| 28788 | 2881  | 
  \end{description}%
 | 
| 26849 | 2882  | 
\end{isamarkuptext}%
 | 
2883  | 
\isamarkuptrue%  | 
|
2884  | 
%  | 
|
2885  | 
\isamarkupsection{Executable code%
 | 
|
2886  | 
}  | 
|
2887  | 
\isamarkuptrue%  | 
|
2888  | 
%  | 
|
2889  | 
\begin{isamarkuptext}%
 | 
|
| 
42627
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2890  | 
For validation purposes, it is often useful to \emph{execute}
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2891  | 
specifications. In principle, execution could be simulated by  | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2892  | 
Isabelle's inference kernel, i.e. by a combination of resolution and  | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2893  | 
simplification. Unfortunately, this approach is rather inefficient.  | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2894  | 
A more efficient way of executing specifications is to translate  | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2895  | 
them into a functional programming language such as ML.  | 
| 26849 | 2896  | 
|
| 45192 | 2897  | 
Isabelle provides a generic framework to support code generation  | 
| 
42627
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2898  | 
from executable specifications. Isabelle/HOL instantiates these  | 
| 45192 | 2899  | 
mechanisms in a way that is amenable to end-user applications. Code  | 
2900  | 
can be generated for functional programs (including overloading  | 
|
2901  | 
  using type classes) targeting SML \cite{SML}, OCaml \cite{OCaml},
 | 
|
2902  | 
  Haskell \cite{haskell-revised-report} and Scala
 | 
|
| 
42627
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2903  | 
  \cite{scala-overview-tech-report}.  Conceptually, code generation is
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2904  | 
  split up in three steps: \emph{selection} of code theorems,
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2905  | 
  \emph{translation} into an abstract executable view and
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2906  | 
  \emph{serialization} to a specific \emph{target language}.
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2907  | 
Inductive specifications can be executed using the predicate  | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2908  | 
  compiler which operates within HOL.  See \cite{isabelle-codegen} for
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
2909  | 
an introduction.  | 
| 37422 | 2910  | 
|
2911  | 
  \begin{matharray}{rcl}
 | 
|
| 40406 | 2912  | 
    \indexdef{HOL}{command}{export\_code}\hypertarget{command.HOL.export-code}{\hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 37422 | 2913  | 
    \indexdef{HOL}{attribute}{code}\hypertarget{attribute.HOL.code}{\hyperlink{attribute.HOL.code}{\mbox{\isa{code}}}} & : & \isa{attribute} \\
 | 
| 40406 | 2914  | 
    \indexdef{HOL}{command}{code\_abort}\hypertarget{command.HOL.code-abort}{\hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}abort}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
2915  | 
    \indexdef{HOL}{command}{code\_datatype}\hypertarget{command.HOL.code-datatype}{\hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2916  | 
    \indexdef{HOL}{command}{print\_codesetup}\hypertarget{command.HOL.print-codesetup}{\hyperlink{command.HOL.print-codesetup}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codesetup}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
2917  | 
    \indexdef{HOL}{attribute}{code\_unfold}\hypertarget{attribute.HOL.code-unfold}{\hyperlink{attribute.HOL.code-unfold}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}unfold}}}} & : & \isa{attribute} \\
 | 
| 40406 | 2918  | 
    \indexdef{HOL}{attribute}{code\_post}\hypertarget{attribute.HOL.code-post}{\hyperlink{attribute.HOL.code-post}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}post}}}} & : & \isa{attribute} \\
 | 
| 46035 | 2919  | 
    \indexdef{HOL}{attribute}{code\_abbrev}\hypertarget{attribute.HOL.code-abbrev}{\hyperlink{attribute.HOL.code-abbrev}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}abbrev}}}} & : & \isa{attribute} \\
 | 
| 40406 | 2920  | 
    \indexdef{HOL}{command}{print\_codeproc}\hypertarget{command.HOL.print-codeproc}{\hyperlink{command.HOL.print-codeproc}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codeproc}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
2921  | 
    \indexdef{HOL}{command}{code\_thms}\hypertarget{command.HOL.code-thms}{\hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}thms}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2922  | 
    \indexdef{HOL}{command}{code\_deps}\hypertarget{command.HOL.code-deps}{\hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}deps}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2923  | 
    \indexdef{HOL}{command}{code\_const}\hypertarget{command.HOL.code-const}{\hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2924  | 
    \indexdef{HOL}{command}{code\_type}\hypertarget{command.HOL.code-type}{\hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2925  | 
    \indexdef{HOL}{command}{code\_class}\hypertarget{command.HOL.code-class}{\hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}class}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2926  | 
    \indexdef{HOL}{command}{code\_instance}\hypertarget{command.HOL.code-instance}{\hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}instance}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2927  | 
    \indexdef{HOL}{command}{code\_reserved}\hypertarget{command.HOL.code-reserved}{\hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reserved}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2928  | 
    \indexdef{HOL}{command}{code\_monad}\hypertarget{command.HOL.code-monad}{\hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}monad}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2929  | 
    \indexdef{HOL}{command}{code\_include}\hypertarget{command.HOL.code-include}{\hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}include}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
2930  | 
    \indexdef{HOL}{command}{code\_modulename}\hypertarget{command.HOL.code-modulename}{\hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}modulename}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 
45408
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
2931  | 
    \indexdef{HOL}{command}{code\_reflect}\hypertarget{command.HOL.code-reflect}{\hyperlink{command.HOL.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
2932  | 
    \indexdef{HOL}{command}{code\_pred}\hypertarget{command.HOL.code-pred}{\hyperlink{command.HOL.code-pred}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}
 | 
| 37422 | 2933  | 
  \end{matharray}
 | 
2934  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2935  | 
  \begin{railoutput}
 | 
| 42662 | 2936  | 
\rail@begin{11}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2937  | 
\rail@term{\hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2938  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2939  | 
\rail@nont{\isa{constexpr}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2940  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2941  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2942  | 
\rail@cr{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2943  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2944  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2945  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2946  | 
\rail@term{\isa{\isakeyword{in}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2947  | 
\rail@nont{\isa{target}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2948  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2949  | 
\rail@nextbar{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2950  | 
\rail@term{\isa{\isakeyword{module{\isaliteral{5F}{\isacharunderscore}}name}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2951  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2952  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2953  | 
\rail@cr{7}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2954  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2955  | 
\rail@nextbar{8}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2956  | 
\rail@term{\isa{\isakeyword{file}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2957  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2958  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2959  | 
\rail@nextbar{9}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2960  | 
\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2961  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2962  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2963  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2964  | 
\rail@nextbar{8}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2965  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2966  | 
\rail@nont{\isa{args}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2967  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2968  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2969  | 
\rail@nextplus{10}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2970  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2971  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2972  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2973  | 
\rail@begin{1}{\isa{const}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2974  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2975  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2976  | 
\rail@begin{3}{\isa{constexpr}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2977  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2978  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2979  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2980  | 
\rail@term{\isa{name{\isaliteral{2E}{\isachardot}}{\isaliteral{5F}{\isacharunderscore}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2981  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2982  | 
\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2983  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2984  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2985  | 
\rail@begin{1}{\isa{typeconstructor}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2986  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2987  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2988  | 
\rail@begin{1}{\isa{class}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2989  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2990  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2991  | 
\rail@begin{4}{\isa{target}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2992  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2993  | 
\rail@term{\isa{SML}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2994  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2995  | 
\rail@term{\isa{OCaml}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2996  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2997  | 
\rail@term{\isa{Haskell}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2998  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2999  | 
\rail@term{\isa{Scala}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3000  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3001  | 
\rail@end  | 
| 42662 | 3002  | 
\rail@begin{4}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3003  | 
\rail@term{\hyperlink{attribute.HOL.code}{\mbox{\isa{code}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3004  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3005  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3006  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3007  | 
\rail@term{\isa{del}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3008  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3009  | 
\rail@term{\isa{abstype}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3010  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3011  | 
\rail@term{\isa{abstract}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3012  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3013  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3014  | 
\rail@end  | 
| 42662 | 3015  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3016  | 
\rail@term{\hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}abort}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3017  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3018  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3019  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3020  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3021  | 
\rail@end  | 
| 42662 | 3022  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3023  | 
\rail@term{\hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}datatype}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3024  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3025  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3026  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3027  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3028  | 
\rail@end  | 
| 42662 | 3029  | 
\rail@begin{2}{}
 | 
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
3030  | 
\rail@term{\hyperlink{attribute.HOL.code-unfold}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}unfold}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3031  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3032  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3033  | 
\rail@term{\isa{del}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3034  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3035  | 
\rail@end  | 
| 42662 | 3036  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3037  | 
\rail@term{\hyperlink{attribute.HOL.code-post}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}post}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3038  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3039  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3040  | 
\rail@term{\isa{del}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3041  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3042  | 
\rail@end  | 
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
3043  | 
\rail@begin{1}{}
 | 
| 46035 | 3044  | 
\rail@term{\hyperlink{attribute.HOL.code-abbrev}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}abbrev}}}}[]
 | 
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
3045  | 
\rail@end  | 
| 42662 | 3046  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3047  | 
\rail@term{\hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}thms}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3048  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3049  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3050  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3051  | 
\rail@nont{\isa{constexpr}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3052  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3053  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3054  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3055  | 
\rail@end  | 
| 42662 | 3056  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3057  | 
\rail@term{\hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}deps}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3058  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3059  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3060  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3061  | 
\rail@nont{\isa{constexpr}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3062  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3063  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3064  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3065  | 
\rail@end  | 
| 42662 | 3066  | 
\rail@begin{7}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3067  | 
\rail@term{\hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3068  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3069  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3070  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3071  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3072  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3073  | 
\rail@cr{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3074  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3075  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3076  | 
\rail@nont{\isa{target}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3077  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3078  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3079  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3080  | 
\rail@nont{\isa{syntax}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3081  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3082  | 
\rail@nextplus{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3083  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3084  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3085  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3086  | 
\rail@nextplus{6}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3087  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3088  | 
\rail@end  | 
| 42662 | 3089  | 
\rail@begin{7}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3090  | 
\rail@term{\hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3091  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3092  | 
\rail@nont{\isa{typeconstructor}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3093  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3094  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3095  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3096  | 
\rail@cr{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3097  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3098  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3099  | 
\rail@nont{\isa{target}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3100  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3101  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3102  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3103  | 
\rail@nont{\isa{syntax}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3104  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3105  | 
\rail@nextplus{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3106  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3107  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3108  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3109  | 
\rail@nextplus{6}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3110  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3111  | 
\rail@end  | 
| 42662 | 3112  | 
\rail@begin{9}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3113  | 
\rail@term{\hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}class}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3114  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3115  | 
\rail@nont{\isa{class}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3116  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3117  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3118  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3119  | 
\rail@cr{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3120  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3121  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3122  | 
\rail@nont{\isa{target}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3123  | 
\rail@cr{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3124  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3125  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3126  | 
\rail@nextbar{6}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3127  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3128  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3129  | 
\rail@nextplus{7}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3130  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3131  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3132  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3133  | 
\rail@nextplus{8}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3134  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3135  | 
\rail@end  | 
| 42662 | 3136  | 
\rail@begin{7}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3137  | 
\rail@term{\hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}instance}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3138  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3139  | 
\rail@nont{\isa{typeconstructor}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3140  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3141  | 
\rail@nont{\isa{class}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3142  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3143  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3144  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3145  | 
\rail@cr{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3146  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3147  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3148  | 
\rail@nont{\isa{target}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3149  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3150  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3151  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3152  | 
\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3153  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3154  | 
\rail@nextplus{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3155  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3156  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3157  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3158  | 
\rail@nextplus{6}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3159  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3160  | 
\rail@end  | 
| 42662 | 3161  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3162  | 
\rail@term{\hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reserved}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3163  | 
\rail@nont{\isa{target}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3164  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3165  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3166  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3167  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3168  | 
\rail@end  | 
| 42662 | 3169  | 
\rail@begin{1}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3170  | 
\rail@term{\hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}monad}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3171  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3172  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3173  | 
\rail@nont{\isa{target}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3174  | 
\rail@end  | 
| 42662 | 3175  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3176  | 
\rail@term{\hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}include}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3177  | 
\rail@nont{\isa{target}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3178  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3179  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3180  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3181  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3182  | 
\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3183  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3184  | 
\rail@end  | 
| 42662 | 3185  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3186  | 
\rail@term{\hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}modulename}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3187  | 
\rail@nont{\isa{target}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3188  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3189  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3190  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3191  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3192  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3193  | 
\rail@end  | 
| 42662 | 3194  | 
\rail@begin{11}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3195  | 
\rail@term{\hyperlink{command.HOL.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3196  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3197  | 
\rail@cr{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3198  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3199  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3200  | 
\rail@term{\isa{\isakeyword{datatypes}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3201  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3202  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3203  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3204  | 
\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3205  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3206  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3207  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3208  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3209  | 
\rail@nextplus{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3210  | 
\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3211  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3212  | 
\rail@nextplus{6}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3213  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3214  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3215  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3216  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3217  | 
\rail@cr{8}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3218  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3219  | 
\rail@nextbar{9}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3220  | 
\rail@term{\isa{\isakeyword{functions}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3221  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3222  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3223  | 
\rail@nextplus{10}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3224  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3225  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3226  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3227  | 
\rail@nextbar{9}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3228  | 
\rail@term{\isa{\isakeyword{file}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3229  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3230  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3231  | 
\rail@end  | 
| 
45408
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3232  | 
\rail@begin{6}{}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3233  | 
\rail@term{\hyperlink{command.HOL.code-pred}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}}}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3234  | 
\rail@cr{2}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3235  | 
\rail@bar  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3236  | 
\rail@nextbar{3}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3237  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3238  | 
\rail@term{\isa{\isakeyword{modes}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3239  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3240  | 
\rail@nont{\isa{modedecl}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3241  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3242  | 
\rail@endbar  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3243  | 
\rail@cr{5}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3244  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3245  | 
\rail@end  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3246  | 
\rail@begin{4}{\isa{syntax}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3247  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3248  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3249  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3250  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3251  | 
\rail@term{\isa{\isakeyword{infix}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3252  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3253  | 
\rail@term{\isa{\isakeyword{infixl}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3254  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3255  | 
\rail@term{\isa{\isakeyword{infixr}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3256  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3257  | 
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3258  | 
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3259  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3260  | 
\rail@end  | 
| 
45408
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3261  | 
\rail@begin{6}{\isa{modedecl}}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3262  | 
\rail@bar  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3263  | 
\rail@nont{\isa{modes}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3264  | 
\rail@nextbar{1}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3265  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3266  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3267  | 
\rail@nont{\isa{modes}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3268  | 
\rail@cr{3}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3269  | 
\rail@bar  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3270  | 
\rail@nextbar{4}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3271  | 
\rail@term{\isa{\isakeyword{and}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3272  | 
\rail@plus  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3273  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3274  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3275  | 
\rail@nont{\isa{modes}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3276  | 
\rail@term{\isa{\isakeyword{and}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3277  | 
\rail@nextplus{5}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3278  | 
\rail@endplus  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3279  | 
\rail@endbar  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3280  | 
\rail@endbar  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3281  | 
\rail@end  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3282  | 
\rail@begin{1}{\isa{modes}}
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3283  | 
\rail@nont{\isa{mode}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3284  | 
\rail@term{\isa{\isakeyword{as}}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3285  | 
\rail@nont{\isa{const}}[]
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3286  | 
\rail@end  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3287  | 
\end{railoutput}
 | 
| 37422 | 3288  | 
|
3289  | 
||
3290  | 
  \begin{description}
 | 
|
3291  | 
||
| 40406 | 3292  | 
  \item \hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}}}} generates code for a given list
 | 
| 39608 | 3293  | 
of constants in the specified target language(s). If no  | 
3294  | 
serialization instruction is given, only abstract code is generated  | 
|
3295  | 
internally.  | 
|
| 37422 | 3296  | 
|
3297  | 
Constants may be specified by giving them literally, referring to  | 
|
| 40406 | 3298  | 
  all executable contants within a certain theory by giving \isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{2E}{\isachardot}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}, or referring to \emph{all} executable constants currently
 | 
3299  | 
  available by giving \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 37422 | 3300  | 
|
3301  | 
By default, for each involved theory one corresponding name space  | 
|
3302  | 
module is generated. Alternativly, a module name may be specified  | 
|
| 40406 | 3303  | 
  after the \hyperlink{keyword.module-name}{\mbox{\isa{\isakeyword{module{\isaliteral{5F}{\isacharunderscore}}name}}}} keyword; then \emph{all} code is
 | 
| 37422 | 3304  | 
placed in this module.  | 
3305  | 
||
| 39608 | 3306  | 
  For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification
 | 
3307  | 
  refers to a single file; for \emph{Haskell}, it refers to a whole
 | 
|
3308  | 
directory, where code is generated in multiple files reflecting the  | 
|
3309  | 
module hierarchy. Omitting the file specification denotes standard  | 
|
| 37749 | 3310  | 
output.  | 
| 37422 | 3311  | 
|
3312  | 
Serializers take an optional list of arguments in parentheses. For  | 
|
| 40406 | 3313  | 
  \emph{SML} and \emph{OCaml}, ``\isa{no{\isaliteral{5F}{\isacharunderscore}}signatures}`` omits
 | 
| 37422 | 3314  | 
explicit module signatures.  | 
| 42123 | 3315  | 
|
| 39608 | 3316  | 
  For \emph{Haskell} a module name prefix may be given using the
 | 
| 40406 | 3317  | 
  ``\isa{{\isaliteral{22}{\isachardoublequote}}root{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}'' argument; ``\isa{string{\isaliteral{5F}{\isacharunderscore}}classes}'' adds a
 | 
| 39608 | 3318  | 
``\verb|deriving (Read, Show)|'' clause to each appropriate  | 
3319  | 
datatype declaration.  | 
|
| 37422 | 3320  | 
|
3321  | 
  \item \hyperlink{attribute.HOL.code}{\mbox{\isa{code}}} explicitly selects (or with option
 | 
|
| 40406 | 3322  | 
  ``\isa{{\isaliteral{22}{\isachardoublequote}}del{\isaliteral{22}{\isachardoublequote}}}'' deselects) a code equation for code generation.
 | 
| 
38462
 
34d3de1254cd
formally document `code abstype` and `code abstract` attributes
 
haftmann 
parents: 
37820 
diff
changeset
 | 
3323  | 
Usually packages introducing code equations provide a reasonable  | 
| 40406 | 3324  | 
  default setup for selection.  Variants \isa{{\isaliteral{22}{\isachardoublequote}}code\ abstype{\isaliteral{22}{\isachardoublequote}}} and
 | 
3325  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}code\ abstract{\isaliteral{22}{\isachardoublequote}}} declare abstract datatype certificates or
 | 
|
| 
38462
 
34d3de1254cd
formally document `code abstype` and `code abstract` attributes
 
haftmann 
parents: 
37820 
diff
changeset
 | 
3326  | 
code equations on abstract datatype representations respectively.  | 
| 37422 | 3327  | 
|
| 40406 | 3328  | 
  \item \hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}abort}}}} declares constants which are not
 | 
| 39608 | 3329  | 
required to have a definition by means of code equations; if needed  | 
3330  | 
these are implemented by program abort instead.  | 
|
| 37422 | 3331  | 
|
| 40406 | 3332  | 
  \item \hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}datatype}}}} specifies a constructor set
 | 
| 37422 | 3333  | 
for a logical type.  | 
3334  | 
||
| 40406 | 3335  | 
  \item \hyperlink{command.HOL.print-codesetup}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codesetup}}}} gives an overview on
 | 
| 37422 | 3336  | 
selected code equations and code generator datatypes.  | 
3337  | 
||
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
3338  | 
  \item \hyperlink{attribute.HOL.code-unfold}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}unfold}}} declares (or with option
 | 
| 46035 | 3339  | 
  ``\isa{{\isaliteral{22}{\isachardoublequote}}del{\isaliteral{22}{\isachardoublequote}}}'' removes) theorems which during preprocessing
 | 
3340  | 
are applied as rewrite rules to any code equation or evaluation  | 
|
3341  | 
input.  | 
|
| 37422 | 3342  | 
|
| 40406 | 3343  | 
  \item \hyperlink{attribute.HOL.code-post}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}post}}} declares (or with option ``\isa{{\isaliteral{22}{\isachardoublequote}}del{\isaliteral{22}{\isachardoublequote}}}'' removes) theorems which are applied as rewrite rules to any
 | 
| 39608 | 3344  | 
result of an evaluation.  | 
| 37422 | 3345  | 
|
| 46035 | 3346  | 
  \item \hyperlink{attribute.HOL.code-abbrev}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}abbrev}}} declares equations which are
 | 
3347  | 
applied as rewrite rules to any result of an evaluation and  | 
|
3348  | 
symmetrically during preprocessing to any code equation or evaluation  | 
|
3349  | 
input.  | 
|
3350  | 
||
| 40406 | 3351  | 
  \item \hyperlink{command.HOL.print-codeproc}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codeproc}}}} prints the setup of the code
 | 
| 39608 | 3352  | 
generator preprocessor.  | 
| 37422 | 3353  | 
|
| 40406 | 3354  | 
  \item \hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}thms}}}} prints a list of theorems
 | 
| 37422 | 3355  | 
representing the corresponding program containing all given  | 
3356  | 
constants after preprocessing.  | 
|
3357  | 
||
| 40406 | 3358  | 
  \item \hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}deps}}}} visualizes dependencies of
 | 
| 37422 | 3359  | 
theorems representing the corresponding program containing all given  | 
3360  | 
constants after preprocessing.  | 
|
3361  | 
||
| 40406 | 3362  | 
  \item \hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}}}} associates a list of constants
 | 
| 37422 | 3363  | 
with target-specific serializations; omitting a serialization  | 
3364  | 
deletes an existing serialization.  | 
|
3365  | 
||
| 40406 | 3366  | 
  \item \hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}}}} associates a list of type
 | 
| 37422 | 3367  | 
constructors with target-specific serializations; omitting a  | 
3368  | 
serialization deletes an existing serialization.  | 
|
3369  | 
||
| 40406 | 3370  | 
  \item \hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}class}}}} associates a list of classes
 | 
| 37422 | 3371  | 
with target-specific class names; omitting a serialization deletes  | 
3372  | 
  an existing serialization.  This applies only to \emph{Haskell}.
 | 
|
3373  | 
||
| 40406 | 3374  | 
  \item \hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}instance}}}} declares a list of type
 | 
| 37422 | 3375  | 
constructor / class instance relations as ``already present'' for a  | 
| 40406 | 3376  | 
  given target.  Omitting a ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{22}{\isachardoublequote}}}'' deletes an existing
 | 
| 37422 | 3377  | 
``already present'' declaration. This applies only to  | 
3378  | 
  \emph{Haskell}.
 | 
|
3379  | 
||
| 40406 | 3380  | 
  \item \hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reserved}}}} declares a list of names as
 | 
| 37422 | 3381  | 
reserved for a given target, preventing it to be shadowed by any  | 
3382  | 
generated code.  | 
|
3383  | 
||
| 40406 | 3384  | 
  \item \hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}monad}}}} provides an auxiliary mechanism
 | 
| 37422 | 3385  | 
to generate monadic code for Haskell.  | 
3386  | 
||
| 40406 | 3387  | 
  \item \hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}include}}}} adds arbitrary named content
 | 
3388  | 
  (``include'') to generated code.  A ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{22}{\isachardoublequote}}}'' as last argument
 | 
|
| 37422 | 3389  | 
will remove an already added ``include''.  | 
3390  | 
||
| 40406 | 3391  | 
  \item \hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}modulename}}}} declares aliasings from one
 | 
| 37422 | 3392  | 
module name onto another.  | 
3393  | 
||
| 40406 | 3394  | 
  \item \hyperlink{command.HOL.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}} without a ``\isa{{\isaliteral{22}{\isachardoublequote}}file{\isaliteral{22}{\isachardoublequote}}}''
 | 
| 39608 | 3395  | 
argument compiles code into the system runtime environment and  | 
3396  | 
modifies the code generator setup that future invocations of system  | 
|
| 47859 | 3397  | 
  runtime code generation referring to one of the ``\isa{{\isaliteral{22}{\isachardoublequote}}datatypes{\isaliteral{22}{\isachardoublequote}}}'' or ``\isa{{\isaliteral{22}{\isachardoublequote}}functions{\isaliteral{22}{\isachardoublequote}}}'' entities use these
 | 
3398  | 
  precompiled entities.  With a ``\isa{{\isaliteral{22}{\isachardoublequote}}file{\isaliteral{22}{\isachardoublequote}}}'' argument, the
 | 
|
3399  | 
corresponding code is generated into that specified file without  | 
|
3400  | 
modifying the code generator setup.  | 
|
3401  | 
||
3402  | 
  \item \hyperlink{command.HOL.code-pred}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}}}} creates code equations for a
 | 
|
3403  | 
predicate given a set of introduction rules. Optional mode  | 
|
3404  | 
annotations determine which arguments are supposed to be input or  | 
|
3405  | 
output. If alternative introduction rules are declared, one must  | 
|
3406  | 
prove a corresponding elimination rule.  | 
|
| 
45408
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
3407  | 
|
| 
42627
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
3408  | 
  \end{description}%
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
3409  | 
\end{isamarkuptext}%
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
3410  | 
\isamarkuptrue%  | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
3411  | 
%  | 
| 27047 | 3412  | 
\isamarkupsection{Definition by specification \label{sec:hol-specification}%
 | 
3413  | 
}  | 
|
3414  | 
\isamarkuptrue%  | 
|
3415  | 
%  | 
|
3416  | 
\begin{isamarkuptext}%
 | 
|
3417  | 
\begin{matharray}{rcl}
 | 
|
| 40406 | 3418  | 
    \indexdef{HOL}{command}{specification}\hypertarget{command.HOL.specification}{\hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
3419  | 
    \indexdef{HOL}{command}{ax\_specification}\hypertarget{command.HOL.ax-specification}{\hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 27047 | 3420  | 
  \end{matharray}
 | 
3421  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3422  | 
  \begin{railoutput}
 | 
| 42662 | 3423  | 
\rail@begin{6}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3424  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3425  | 
\rail@term{\hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3426  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3427  | 
\rail@term{\hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3428  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3429  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3430  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3431  | 
\rail@nont{\isa{decl}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3432  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3433  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3434  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3435  | 
\rail@cr{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3436  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3437  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3438  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3439  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3440  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3441  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3442  | 
\rail@nextplus{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3443  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3444  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3445  | 
\rail@begin{2}{\isa{decl}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3446  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3447  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3448  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3449  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3450  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3451  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3452  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3453  | 
\rail@term{\isa{\isakeyword{overloaded}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3454  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3455  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3456  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3457  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3458  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3459  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
3460  | 
|
| 27047 | 3461  | 
|
| 28788 | 3462  | 
  \begin{description}
 | 
| 27047 | 3463  | 
|
| 40406 | 3464  | 
  \item \hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}decls\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} sets up a
 | 
| 27047 | 3465  | 
goal stating the existence of terms with the properties specified to  | 
3466  | 
  hold for the constants given in \isa{decls}.  After finishing the
 | 
|
3467  | 
proof, the theory will be augmented with definitions for the given  | 
|
3468  | 
constants, as well as with theorems stating the properties for these  | 
|
3469  | 
constants.  | 
|
3470  | 
||
| 40406 | 3471  | 
  \item \hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}decls\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} sets up
 | 
| 28788 | 3472  | 
a goal stating the existence of terms with the properties specified  | 
3473  | 
  to hold for the constants given in \isa{decls}.  After finishing
 | 
|
3474  | 
the proof, the theory will be augmented with axioms expressing the  | 
|
3475  | 
properties given in the first place.  | 
|
| 27047 | 3476  | 
|
| 28788 | 3477  | 
  \item \isa{decl} declares a constant to be defined by the
 | 
| 27047 | 3478  | 
  specification given.  The definition for the constant \isa{c} is
 | 
| 40406 | 3479  | 
  bound to the name \isa{c{\isaliteral{5F}{\isacharunderscore}}def} unless a theorem name is given in
 | 
| 27047 | 3480  | 
the declaration. Overloaded constants should be declared as such.  | 
3481  | 
||
| 28788 | 3482  | 
  \end{description}
 | 
| 27047 | 3483  | 
|
| 40406 | 3484  | 
  Whether to use \hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}} or \hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}} is to some extent a matter of style.  \hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}} introduces no new axioms, and so by
 | 
3485  | 
  construction cannot introduce inconsistencies, whereas \hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}} does introduce axioms, but only after the
 | 
|
| 27047 | 3486  | 
user has explicitly proven it to be safe. A practical issue must be  | 
3487  | 
considered, though: After introducing two constants with the same  | 
|
3488  | 
  properties using \hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}}, one can prove
 | 
|
3489  | 
that the two constants are, in fact, equal. If this might be a  | 
|
| 40406 | 3490  | 
  problem, one should use \hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}}.%
 | 
| 27047 | 3491  | 
\end{isamarkuptext}%
 | 
3492  | 
\isamarkuptrue%  | 
|
3493  | 
%  | 
|
| 26849 | 3494  | 
\isadelimtheory  | 
3495  | 
%  | 
|
3496  | 
\endisadelimtheory  | 
|
3497  | 
%  | 
|
3498  | 
\isatagtheory  | 
|
| 26840 | 3499  | 
\isacommand{end}\isamarkupfalse%
 | 
3500  | 
%  | 
|
3501  | 
\endisatagtheory  | 
|
3502  | 
{\isafoldtheory}%
 | 
|
3503  | 
%  | 
|
3504  | 
\isadelimtheory  | 
|
3505  | 
%  | 
|
3506  | 
\endisadelimtheory  | 
|
| 26849 | 3507  | 
\isanewline  | 
| 26840 | 3508  | 
\end{isabellebody}%
 | 
3509  | 
%%% Local Variables:  | 
|
3510  | 
%%% mode: latex  | 
|
3511  | 
%%% TeX-master: "root"  | 
|
3512  | 
%%% End:  |