| author | wenzelm | 
| Fri, 11 Nov 2011 16:06:26 +0100 | |
| changeset 45457 | 615ba724b269 | 
| parent 45409 | 5abb0e738b00 | 
| child 45678 | 1a6206f538d4 | 
| permissions | -rw-r--r-- | 
| 26840 | 1  | 
theory HOL_Specific  | 
| 44055 | 2  | 
imports Base Main "~~/src/HOL/Library/Old_Recdef"  | 
| 26840 | 3  | 
begin  | 
4  | 
||
| 26852 | 5  | 
chapter {* Isabelle/HOL \label{ch:hol} *}
 | 
| 26849 | 6  | 
|
| 42914 | 7  | 
section {* Higher-Order Logic *}
 | 
8  | 
||
9  | 
text {* Isabelle/HOL is based on Higher-Order Logic, a polymorphic
 | 
|
10  | 
version of Church's Simple Theory of Types. HOL can be best  | 
|
11  | 
understood as a simply-typed version of classical set theory. The  | 
|
12  | 
logic was first implemented in Gordon's HOL system  | 
|
13  | 
  \cite{mgordon-hol}.  It extends Church's original logic
 | 
|
14  | 
  \cite{church40} by explicit type variables (naive polymorphism) and
 | 
|
15  | 
a sound axiomatization scheme for new types based on subsets of  | 
|
16  | 
existing types.  | 
|
17  | 
||
18  | 
  Andrews's book \cite{andrews86} is a full description of the
 | 
|
19  | 
original Church-style higher-order logic, with proofs of correctness  | 
|
20  | 
and completeness wrt.\ certain set-theoretic interpretations. The  | 
|
21  | 
particular extensions of Gordon-style HOL are explained semantically  | 
|
22  | 
  in two chapters of the 1993 HOL book \cite{pitts93}.
 | 
|
23  | 
||
24  | 
Experience with HOL over decades has demonstrated that higher-order  | 
|
25  | 
logic is widely applicable in many areas of mathematics and computer  | 
|
26  | 
science. In a sense, Higher-Order Logic is simpler than First-Order  | 
|
27  | 
Logic, because there are fewer restrictions and special cases. Note  | 
|
28  | 
  that HOL is \emph{weaker} than FOL with axioms for ZF set theory,
 | 
|
29  | 
which is traditionally considered the standard foundation of regular  | 
|
30  | 
mathematics, but for most applications this does not matter. If you  | 
|
31  | 
prefer ML to Lisp, you will probably prefer HOL to ZF.  | 
|
32  | 
||
33  | 
  \medskip The syntax of HOL follows @{text "\<lambda>"}-calculus and
 | 
|
34  | 
functional programming. Function application is curried. To apply  | 
|
35  | 
  the function @{text f} of type @{text "\<tau>\<^sub>1 \<Rightarrow> \<tau>\<^sub>2 \<Rightarrow> \<tau>\<^sub>3"} to the
 | 
|
36  | 
  arguments @{text a} and @{text b} in HOL, you simply write @{text "f
 | 
|
37  | 
a b"} (as in ML or Haskell). There is no ``apply'' operator; the  | 
|
38  | 
  existing application of the Pure @{text "\<lambda>"}-calculus is re-used.
 | 
|
39  | 
  Note that in HOL @{text "f (a, b)"} means ``@{text "f"} applied to
 | 
|
40  | 
  the pair @{text "(a, b)"} (which is notation for @{text "Pair a
 | 
|
41  | 
b"}). The latter typically introduces extra formal efforts that can  | 
|
42  | 
be avoided by currying functions by default. Explicit tuples are as  | 
|
43  | 
infrequent in HOL formalizations as in good ML or Haskell programs.  | 
|
44  | 
||
45  | 
\medskip Isabelle/HOL has a distinct feel, compared to other  | 
|
46  | 
object-logics like Isabelle/ZF. It identifies object-level types  | 
|
47  | 
with meta-level types, taking advantage of the default  | 
|
48  | 
type-inference mechanism of Isabelle/Pure. HOL fully identifies  | 
|
49  | 
object-level functions with meta-level functions, with native  | 
|
50  | 
abstraction and application.  | 
|
51  | 
||
52  | 
These identifications allow Isabelle to support HOL particularly  | 
|
53  | 
nicely, but they also mean that HOL requires some sophistication  | 
|
54  | 
from the user. In particular, an understanding of Hindley-Milner  | 
|
55  | 
type-inference with type-classes, which are both used extensively in  | 
|
56  | 
the standard libraries and applications. Beginners can set  | 
|
57  | 
  @{attribute show_types} or even @{attribute show_sorts} to get more
 | 
|
58  | 
explicit information about the result of type-inference. *}  | 
|
59  | 
||
60  | 
||
| 42908 | 61  | 
section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
 | 
62  | 
||
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
63  | 
text {* An \emph{inductive definition} specifies the least predicate
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
64  | 
  or set @{text R} closed under given rules: applying a rule to
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
65  | 
  elements of @{text R} yields a result within @{text R}.  For
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
66  | 
example, a structural operational semantics is an inductive  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
67  | 
definition of an evaluation relation.  | 
| 42908 | 68  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
69  | 
  Dually, a \emph{coinductive definition} specifies the greatest
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
70  | 
  predicate or set @{text R} that is consistent with given rules:
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
71  | 
  every element of @{text 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
 | 
72  | 
  to elements of @{text R}.  An important example is using
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
73  | 
bisimulation relations to formalise equivalence of processes and  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
74  | 
infinite data structures.  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
75  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
76  | 
Both inductive and coinductive definitions are based on the  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
77  | 
Knaster-Tarski fixed-point theorem for complete lattices. The  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
78  | 
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
 | 
79  | 
functor on subsets of set-theoretic relations. The required  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
80  | 
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
 | 
81  | 
the fixed-point definition and the resulting consequences. This  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
82  | 
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
 | 
83  | 
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
 | 
84  | 
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
 | 
85  | 
  @{text "A \<le> B \<Longrightarrow> \<M> A \<le> \<M> B"}, for each premise @{text "\<M> R t"} in an
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
86  | 
introduction rule. The default rule declarations of Isabelle/HOL  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
87  | 
already take care of most common situations.  | 
| 
42907
 
dfd4ef8e73f6
updated and re-unified HOL typedef, with some live examples;
 
wenzelm 
parents: 
42705 
diff
changeset
 | 
88  | 
|
| 26849 | 89  | 
  \begin{matharray}{rcl}
 | 
| 42908 | 90  | 
    @{command_def (HOL) "inductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
91  | 
    @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
|
92  | 
    @{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
|
93  | 
    @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
|
94  | 
    @{attribute_def (HOL) mono} & : & @{text attribute} \\
 | 
|
| 26849 | 95  | 
  \end{matharray}
 | 
96  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
97  | 
  @{rail "
 | 
| 42908 | 98  | 
    (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} |
 | 
99  | 
      @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set})
 | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
100  | 
    @{syntax target}? \\
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
101  | 
    @{syntax \"fixes\"} (@'for' @{syntax \"fixes\"})? (@'where' clauses)? \\
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
102  | 
    (@'monos' @{syntax thmrefs})?
 | 
| 42908 | 103  | 
;  | 
104  | 
    clauses: (@{syntax thmdecl}? @{syntax prop} + '|')
 | 
|
105  | 
;  | 
|
106  | 
    @@{attribute (HOL) mono} (() | 'add' | 'del')
 | 
|
107  | 
"}  | 
|
108  | 
||
109  | 
  \begin{description}
 | 
|
110  | 
||
111  | 
  \item @{command (HOL) "inductive"} and @{command (HOL)
 | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
112  | 
"coinductive"} define (co)inductive predicates from the introduction  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
113  | 
rules.  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
114  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
115  | 
  The propositions given as @{text "clauses"} in the @{keyword
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
116  | 
  "where"} part are either rules of the usual @{text "\<And>/\<Longrightarrow>"} format
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
117  | 
  (with arbitrary nesting), or equalities using @{text "\<equiv>"}.  The
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
118  | 
latter specifies extra-logical abbreviations in the sense of  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
119  | 
  @{command_ref abbreviation}.  Introducing abstract syntax
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
120  | 
simultaneously with the actual introduction rules is occasionally  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
121  | 
useful for complex specifications.  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
122  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
123  | 
  The optional @{keyword "for"} part contains a list of parameters of
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
124  | 
the (co)inductive predicates that remain fixed throughout the  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
125  | 
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
 | 
126  | 
  in each occurrence within the given @{text "clauses"}.
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
127  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
128  | 
  The optional @{keyword "monos"} declaration contains additional
 | 
| 42908 | 129  | 
  \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
 | 
130  | 
applied to a recursive set in the introduction rules.  | 
| 42908 | 131  | 
|
132  | 
  \item @{command (HOL) "inductive_set"} and @{command (HOL)
 | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
133  | 
"coinductive_set"} are wrappers for to the previous commands for  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
134  | 
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
 | 
135  | 
where multiple arguments are simulated via tuples.  | 
| 42908 | 136  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
137  | 
  \item @{attribute (HOL) mono} declares monotonicity rules in the
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
138  | 
context. These rule are involved in the automated monotonicity  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
139  | 
proof of the above inductive and coinductive definitions.  | 
| 42908 | 140  | 
|
141  | 
  \end{description}
 | 
|
142  | 
*}  | 
|
143  | 
||
144  | 
||
145  | 
subsection {* Derived rules *}
 | 
|
146  | 
||
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
147  | 
text {* A (co)inductive definition of @{text R} provides the following
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
148  | 
main theorems:  | 
| 42908 | 149  | 
|
150  | 
  \begin{description}
 | 
|
151  | 
||
152  | 
  \item @{text R.intros} is the list of introduction rules as proven
 | 
|
153  | 
theorems, for the recursive predicates (or sets). The rules are  | 
|
154  | 
also available individually, using the names given them in the  | 
|
155  | 
theory file;  | 
|
156  | 
||
157  | 
  \item @{text R.cases} is the case analysis (or elimination) rule;
 | 
|
158  | 
||
159  | 
  \item @{text R.induct} or @{text R.coinduct} is the (co)induction
 | 
|
| 
44273
 
336752fb25df
adding documentation about simps equation in the inductive package
 
bulwahn 
parents: 
44055 
diff
changeset
 | 
160  | 
rule;  | 
| 
 
336752fb25df
adding documentation about simps equation in the inductive package
 
bulwahn 
parents: 
44055 
diff
changeset
 | 
161  | 
|
| 
 
336752fb25df
adding documentation about simps equation in the inductive package
 
bulwahn 
parents: 
44055 
diff
changeset
 | 
162  | 
  \item @{text R.simps} is the equation unrolling the fixpoint of the
 | 
| 
 
336752fb25df
adding documentation about simps equation in the inductive package
 
bulwahn 
parents: 
44055 
diff
changeset
 | 
163  | 
predicate one step.  | 
| 
 
336752fb25df
adding documentation about simps equation in the inductive package
 
bulwahn 
parents: 
44055 
diff
changeset
 | 
164  | 
|
| 42908 | 165  | 
  \end{description}
 | 
166  | 
||
167  | 
  When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
 | 
|
168  | 
defined simultaneously, the list of introduction rules is called  | 
|
169  | 
  @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
 | 
|
170  | 
  called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
 | 
|
171  | 
  of mutual induction rules is called @{text
 | 
|
172  | 
"R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.  | 
|
173  | 
*}  | 
|
174  | 
||
175  | 
||
176  | 
subsection {* Monotonicity theorems *}
 | 
|
177  | 
||
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
178  | 
text {* 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
 | 
179  | 
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
 | 
180  | 
  @{attribute (HOL) mono} attribute.  See the main Isabelle/HOL
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
181  | 
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
 | 
182  | 
theorems is as follows:  | 
| 42908 | 183  | 
|
184  | 
  \begin{itemize}
 | 
|
185  | 
||
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
186  | 
  \item Theorems of the form @{text "A \<le> B \<Longrightarrow> \<M> A \<le> \<M> B"}, for proving
 | 
| 42908 | 187  | 
monotonicity of inductive definitions whose introduction rules have  | 
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
188  | 
  premises involving terms such as @{text "\<M> R t"}.
 | 
| 42908 | 189  | 
|
190  | 
\item Monotonicity theorems for logical operators, which are of the  | 
|
191  | 
  general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}.  For example, in
 | 
|
192  | 
  the case of the operator @{text "\<or>"}, the corresponding theorem is
 | 
|
193  | 
\[  | 
|
194  | 
  \infer{@{text "P\<^sub>1 \<or> P\<^sub>2 \<longrightarrow> Q\<^sub>1 \<or> Q\<^sub>2"}}{@{text "P\<^sub>1 \<longrightarrow> Q\<^sub>1"} & @{text "P\<^sub>2 \<longrightarrow> Q\<^sub>2"}}
 | 
|
195  | 
\]  | 
|
196  | 
||
197  | 
\item De Morgan style equations for reasoning about the ``polarity''  | 
|
198  | 
of expressions, e.g.  | 
|
199  | 
\[  | 
|
200  | 
  @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
 | 
|
201  | 
  @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
 | 
|
202  | 
\]  | 
|
203  | 
||
204  | 
\item Equations for reducing complex operators to more primitive  | 
|
205  | 
ones whose monotonicity can easily be proved, e.g.  | 
|
206  | 
\[  | 
|
207  | 
  @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
 | 
|
208  | 
  @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
 | 
|
209  | 
\]  | 
|
210  | 
||
211  | 
  \end{itemize}
 | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
212  | 
*}  | 
| 42908 | 213  | 
|
| 
42913
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
214  | 
subsubsection {* Examples *}
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
215  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
216  | 
text {* 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
 | 
217  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
218  | 
inductive_set Fin :: "'a set \<Rightarrow> 'a set set" for A :: "'a set"  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
219  | 
where  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
220  | 
  empty: "{} \<in> Fin A"
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
221  | 
| insert: "a \<in> A \<Longrightarrow> B \<in> Fin A \<Longrightarrow> insert a B \<in> Fin A"  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
222  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
223  | 
text {* 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
 | 
224  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
225  | 
inductive acc :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
226  | 
for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<prec>" 50)  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
227  | 
where acc: "(\<And>y. y \<prec> x \<Longrightarrow> acc r y) \<Longrightarrow> acc r x"  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
228  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
229  | 
text {* Common logical connectives can be easily characterized as
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
230  | 
non-recursive inductive definitions with parameters, but without  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
231  | 
arguments. *}  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
232  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
233  | 
inductive AND for A B :: bool  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
234  | 
where "A \<Longrightarrow> B \<Longrightarrow> AND A B"  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
235  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
236  | 
inductive OR for A B :: bool  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
237  | 
where "A \<Longrightarrow> OR A B"  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
238  | 
| "B \<Longrightarrow> OR A B"  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
239  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
240  | 
inductive EXISTS for B :: "'a \<Rightarrow> bool"  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
241  | 
where "B a \<Longrightarrow> EXISTS B"  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
242  | 
|
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
243  | 
text {* Here the @{text "cases"} or @{text "induct"} rules produced by
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
244  | 
  the @{command inductive} package coincide with the expected
 | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
245  | 
elimination rules for Natural Deduction. Already in the original  | 
| 
 
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
 
wenzelm 
parents: 
42912 
diff
changeset
 | 
246  | 
  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
 | 
247  | 
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
 | 
248  | 
elimination can be constructed systematically. *}  | 
| 42908 | 249  | 
|
250  | 
||
251  | 
section {* Recursive functions \label{sec:recursion} *}
 | 
|
252  | 
||
253  | 
text {*
 | 
|
254  | 
  \begin{matharray}{rcl}
 | 
|
255  | 
    @{command_def (HOL) "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
|
256  | 
    @{command_def (HOL) "fun"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
|
257  | 
    @{command_def (HOL) "function"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
 | 
|
258  | 
    @{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
 | 
|
259  | 
  \end{matharray}
 | 
|
260  | 
||
261  | 
  @{rail "
 | 
|
262  | 
    @@{command (HOL) primrec} @{syntax target}? @{syntax \"fixes\"} @'where' equations
 | 
|
263  | 
;  | 
|
264  | 
    (@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts?
 | 
|
265  | 
      @{syntax \"fixes\"} \\ @'where' equations
 | 
|
| 26849 | 266  | 
;  | 
267  | 
||
| 42908 | 268  | 
    equations: (@{syntax thmdecl}? @{syntax prop} + '|')
 | 
| 26849 | 269  | 
;  | 
| 42908 | 270  | 
    functionopts: '(' (('sequential' | 'domintros') + ',') ')'
 | 
| 26849 | 271  | 
;  | 
| 42908 | 272  | 
    @@{command (HOL) termination} @{syntax term}?
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
273  | 
"}  | 
| 26849 | 274  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
275  | 
  \begin{description}
 | 
| 42123 | 276  | 
|
| 42908 | 277  | 
  \item @{command (HOL) "primrec"} defines primitive recursive
 | 
| 42912 | 278  | 
  functions over datatypes (see also @{command_ref (HOL) datatype} and
 | 
279  | 
  @{command_ref (HOL) rep_datatype}).  The given @{text equations}
 | 
|
280  | 
specify reduction rules that are produced by instantiating the  | 
|
281  | 
generic combinator for primitive recursion that is available for  | 
|
282  | 
each datatype.  | 
|
283  | 
||
284  | 
Each equation needs to be of the form:  | 
|
285  | 
||
286  | 
  @{text [display] "f x\<^sub>1 \<dots> x\<^sub>m (C y\<^sub>1 \<dots> y\<^sub>k) z\<^sub>1 \<dots> z\<^sub>n = rhs"}
 | 
|
287  | 
||
288  | 
  such that @{text C} is a datatype constructor, @{text rhs} contains
 | 
|
289  | 
only the free variables on the left-hand side (or from the context),  | 
|
290  | 
  and all recursive occurrences of @{text "f"} in @{text "rhs"} are of
 | 
|
291  | 
  the form @{text "f \<dots> y\<^sub>i \<dots>"} for some @{text i}.  At most one
 | 
|
292  | 
reduction rule for each constructor can be given. The order does  | 
|
293  | 
not matter. For missing constructors, the function is defined to  | 
|
294  | 
return a default value, but this equation is made difficult to  | 
|
295  | 
access for users.  | 
|
296  | 
||
297  | 
  The reduction rules are declared as @{attribute simp} by default,
 | 
|
298  | 
  which enables standard proof methods like @{method simp} and
 | 
|
299  | 
  @{method auto} to normalize expressions of @{text "f"} applied to
 | 
|
300  | 
datatype constructions, by simulating symbolic computation via  | 
|
301  | 
rewriting.  | 
|
| 35744 | 302  | 
|
| 42908 | 303  | 
  \item @{command (HOL) "function"} defines functions by general
 | 
304  | 
wellfounded recursion. A detailed description with examples can be  | 
|
305  | 
  found in \cite{isabelle-function}. The function is specified by a
 | 
|
306  | 
set of (possibly conditional) recursive equations with arbitrary  | 
|
307  | 
pattern matching. The command generates proof obligations for the  | 
|
308  | 
completeness and the compatibility of patterns.  | 
|
| 
42907
 
dfd4ef8e73f6
updated and re-unified HOL typedef, with some live examples;
 
wenzelm 
parents: 
42705 
diff
changeset
 | 
309  | 
|
| 42908 | 310  | 
The defined function is considered partial, and the resulting  | 
311  | 
  simplification rules (named @{text "f.psimps"}) and induction rule
 | 
|
312  | 
  (named @{text "f.pinduct"}) are guarded by a generated domain
 | 
|
313  | 
  predicate @{text "f_dom"}. The @{command (HOL) "termination"}
 | 
|
314  | 
command can then be used to establish that the function is total.  | 
|
| 42123 | 315  | 
|
| 42908 | 316  | 
  \item @{command (HOL) "fun"} is a shorthand notation for ``@{command
 | 
317  | 
  (HOL) "function"}~@{text "(sequential)"}, followed by automated
 | 
|
318  | 
proof attempts regarding pattern matching and termination. See  | 
|
319  | 
  \cite{isabelle-function} for further details.
 | 
|
| 42123 | 320  | 
|
| 42908 | 321  | 
  \item @{command (HOL) "termination"}~@{text f} commences a
 | 
322  | 
  termination proof for the previously defined function @{text f}.  If
 | 
|
323  | 
this is omitted, the command refers to the most recent function  | 
|
324  | 
definition. After the proof is closed, the recursive equations and  | 
|
325  | 
the induction principle is established.  | 
|
| 26849 | 326  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
327  | 
  \end{description}
 | 
| 
42907
 
dfd4ef8e73f6
updated and re-unified HOL typedef, with some live examples;
 
wenzelm 
parents: 
42705 
diff
changeset
 | 
328  | 
|
| 42908 | 329  | 
  Recursive definitions introduced by the @{command (HOL) "function"}
 | 
| 42912 | 330  | 
  command accommodate reasoning by induction (cf.\ @{method induct}):
 | 
331  | 
  rule @{text "f.induct"} refers to a specific induction rule, with
 | 
|
332  | 
parameters named according to the user-specified equations. Cases  | 
|
333  | 
  are numbered starting from 1.  For @{command (HOL) "primrec"}, the
 | 
|
334  | 
induction principle coincides with structural recursion on the  | 
|
335  | 
datatype where the recursion is carried out.  | 
|
| 42908 | 336  | 
|
337  | 
The equations provided by these packages may be referred later as  | 
|
338  | 
  theorem list @{text "f.simps"}, where @{text f} is the (collective)
 | 
|
339  | 
name of the functions defined. Individual equations may be named  | 
|
340  | 
explicitly as well.  | 
|
341  | 
||
342  | 
  The @{command (HOL) "function"} command accepts the following
 | 
|
343  | 
options.  | 
|
344  | 
||
345  | 
  \begin{description}
 | 
|
346  | 
||
347  | 
  \item @{text sequential} enables a preprocessor which disambiguates
 | 
|
348  | 
overlapping patterns by making them mutually disjoint. Earlier  | 
|
349  | 
equations take precedence over later ones. This allows to give the  | 
|
350  | 
specification in a format very similar to functional programming.  | 
|
351  | 
Note that the resulting simplification and induction rules  | 
|
352  | 
correspond to the transformed specification, not the one given  | 
|
353  | 
originally. This usually means that each equation given by the user  | 
|
354  | 
may result in several theorems. Also note that this automatic  | 
|
355  | 
transformation only works for ML-style datatype patterns.  | 
|
356  | 
||
357  | 
  \item @{text domintros} enables the automated generation of
 | 
|
358  | 
introduction rules for the domain predicate. While mostly not  | 
|
359  | 
needed, they can be helpful in some proofs about partial functions.  | 
|
360  | 
||
361  | 
  \end{description}
 | 
|
| 26849 | 362  | 
*}  | 
363  | 
||
| 42912 | 364  | 
subsubsection {* Example: evaluation of expressions *}
 | 
365  | 
||
366  | 
text {* Subsequently, we define mutual datatypes for arithmetic and
 | 
|
367  | 
  boolean expressions, and use @{command primrec} for evaluation
 | 
|
368  | 
functions that follow the same recursive structure. *}  | 
|
369  | 
||
370  | 
datatype 'a aexp =  | 
|
371  | 
IF "'a bexp" "'a aexp" "'a aexp"  | 
|
372  | 
| Sum "'a aexp" "'a aexp"  | 
|
373  | 
| Diff "'a aexp" "'a aexp"  | 
|
374  | 
| Var 'a  | 
|
375  | 
| Num nat  | 
|
376  | 
and 'a bexp =  | 
|
377  | 
Less "'a aexp" "'a aexp"  | 
|
378  | 
| And "'a bexp" "'a bexp"  | 
|
379  | 
| Neg "'a bexp"  | 
|
380  | 
||
381  | 
||
382  | 
text {* \medskip Evaluation of arithmetic and boolean expressions *}
 | 
|
383  | 
||
384  | 
primrec evala :: "('a \<Rightarrow> nat) \<Rightarrow> 'a aexp \<Rightarrow> nat"
 | 
|
385  | 
  and evalb :: "('a \<Rightarrow> nat) \<Rightarrow> 'a bexp \<Rightarrow> bool"
 | 
|
386  | 
where  | 
|
387  | 
"evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)"  | 
|
388  | 
| "evala env (Sum a1 a2) = evala env a1 + evala env a2"  | 
|
389  | 
| "evala env (Diff a1 a2) = evala env a1 - evala env a2"  | 
|
390  | 
| "evala env (Var v) = env v"  | 
|
391  | 
| "evala env (Num n) = n"  | 
|
392  | 
| "evalb env (Less a1 a2) = (evala env a1 < evala env a2)"  | 
|
393  | 
| "evalb env (And b1 b2) = (evalb env b1 \<and> evalb env b2)"  | 
|
394  | 
| "evalb env (Neg b) = (\<not> evalb env b)"  | 
|
395  | 
||
396  | 
text {* Since the value of an expression depends on the value of its
 | 
|
397  | 
  variables, the functions @{const evala} and @{const evalb} take an
 | 
|
398  | 
  additional parameter, an \emph{environment} that maps variables to
 | 
|
399  | 
their values.  | 
|
400  | 
||
401  | 
\medskip Substitution on expressions can be defined similarly. The  | 
|
402  | 
  mapping @{text f} of type @{typ "'a \<Rightarrow> 'a aexp"} given as a
 | 
|
403  | 
  parameter is lifted canonically on the types @{typ "'a aexp"} and
 | 
|
404  | 
  @{typ "'a bexp"}, respectively.
 | 
|
405  | 
*}  | 
|
406  | 
||
407  | 
primrec substa :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a aexp \<Rightarrow> 'b aexp"
 | 
|
408  | 
  and substb :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a bexp \<Rightarrow> 'b bexp"
 | 
|
409  | 
where  | 
|
410  | 
"substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)"  | 
|
411  | 
| "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)"  | 
|
412  | 
| "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)"  | 
|
413  | 
| "substa f (Var v) = f v"  | 
|
414  | 
| "substa f (Num n) = Num n"  | 
|
415  | 
| "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)"  | 
|
416  | 
| "substb f (And b1 b2) = And (substb f b1) (substb f b2)"  | 
|
417  | 
| "substb f (Neg b) = Neg (substb f b)"  | 
|
418  | 
||
419  | 
text {* In textbooks about semantics one often finds substitution
 | 
|
420  | 
theorems, which express the relationship between substitution and  | 
|
421  | 
  evaluation.  For @{typ "'a aexp"} and @{typ "'a bexp"}, we can prove
 | 
|
422  | 
such a theorem by mutual induction, followed by simplification.  | 
|
423  | 
*}  | 
|
424  | 
||
425  | 
lemma subst_one:  | 
|
426  | 
"evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a"  | 
|
427  | 
"evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b"  | 
|
428  | 
by (induct a and b) simp_all  | 
|
429  | 
||
430  | 
lemma subst_all:  | 
|
431  | 
"evala env (substa s a) = evala (\<lambda>x. evala env (s x)) a"  | 
|
432  | 
"evalb env (substb s b) = evalb (\<lambda>x. evala env (s x)) b"  | 
|
433  | 
by (induct a and b) simp_all  | 
|
434  | 
||
435  | 
||
436  | 
subsubsection {* Example: a substitution function for terms *}
 | 
|
437  | 
||
438  | 
text {* Functions on datatypes with nested recursion are also defined
 | 
|
439  | 
by mutual primitive recursion. *}  | 
|
440  | 
||
441  | 
datatype ('a, 'b) "term" = Var 'a | App 'b "('a, 'b) term list"
 | 
|
442  | 
||
443  | 
text {* A substitution function on type @{typ "('a, 'b) term"} can be
 | 
|
444  | 
  defined as follows, by working simultaneously on @{typ "('a, 'b)
 | 
|
445  | 
term list"}: *}  | 
|
446  | 
||
447  | 
primrec subst_term :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term \<Rightarrow> ('a, 'b) term" and
 | 
|
448  | 
  subst_term_list :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term list \<Rightarrow> ('a, 'b) term list"
 | 
|
449  | 
where  | 
|
450  | 
"subst_term f (Var a) = f a"  | 
|
451  | 
| "subst_term f (App b ts) = App b (subst_term_list f ts)"  | 
|
452  | 
| "subst_term_list f [] = []"  | 
|
453  | 
| "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts"  | 
|
454  | 
||
455  | 
text {* The recursion scheme follows the structure of the unfolded
 | 
|
456  | 
  definition of type @{typ "('a, 'b) term"}.  To prove properties of this
 | 
|
457  | 
substitution function, mutual induction is needed:  | 
|
458  | 
*}  | 
|
459  | 
||
460  | 
lemma "subst_term (subst_term f1 \<circ> f2) t = subst_term f1 (subst_term f2 t)" and  | 
|
461  | 
"subst_term_list (subst_term f1 \<circ> f2) ts = subst_term_list f1 (subst_term_list f2 ts)"  | 
|
462  | 
by (induct t and ts) simp_all  | 
|
463  | 
||
464  | 
||
465  | 
subsubsection {* Example: a map function for infinitely branching trees *}
 | 
|
466  | 
||
467  | 
text {* Defining functions on infinitely branching datatypes by
 | 
|
468  | 
primitive recursion is just as easy.  | 
|
469  | 
*}  | 
|
470  | 
||
471  | 
datatype 'a tree = Atom 'a | Branch "nat \<Rightarrow> 'a tree"  | 
|
472  | 
||
473  | 
primrec map_tree :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a tree \<Rightarrow> 'b tree"
 | 
|
474  | 
where  | 
|
475  | 
"map_tree f (Atom a) = Atom (f a)"  | 
|
476  | 
| "map_tree f (Branch ts) = Branch (\<lambda>x. map_tree f (ts x))"  | 
|
477  | 
||
478  | 
text {* Note that all occurrences of functions such as @{text ts}
 | 
|
479  | 
  above must be applied to an argument.  In particular, @{term
 | 
|
480  | 
"map_tree f \<circ> ts"} is not allowed here. *}  | 
|
481  | 
||
482  | 
text {* Here is a simple composition lemma for @{term map_tree}: *}
 | 
|
483  | 
||
484  | 
lemma "map_tree g (map_tree f t) = map_tree (g \<circ> f) t"  | 
|
485  | 
by (induct t) simp_all  | 
|
486  | 
||
| 
42907
 
dfd4ef8e73f6
updated and re-unified HOL typedef, with some live examples;
 
wenzelm 
parents: 
42705 
diff
changeset
 | 
487  | 
|
| 42908 | 488  | 
subsection {* Proof methods related to recursive definitions *}
 | 
| 26849 | 489  | 
|
490  | 
text {*
 | 
|
491  | 
  \begin{matharray}{rcl}
 | 
|
| 42908 | 492  | 
    @{method_def (HOL) pat_completeness} & : & @{text method} \\
 | 
493  | 
    @{method_def (HOL) relation} & : & @{text method} \\
 | 
|
494  | 
    @{method_def (HOL) lexicographic_order} & : & @{text method} \\
 | 
|
495  | 
    @{method_def (HOL) size_change} & : & @{text method} \\
 | 
|
| 26849 | 496  | 
  \end{matharray}
 | 
497  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
498  | 
  @{rail "
 | 
| 42908 | 499  | 
    @@{method (HOL) relation} @{syntax term}
 | 
500  | 
;  | 
|
501  | 
    @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * )
 | 
|
502  | 
;  | 
|
503  | 
    @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) )
 | 
|
504  | 
;  | 
|
505  | 
orders: ( 'max' | 'min' | 'ms' ) *  | 
|
506  | 
"}  | 
|
507  | 
||
508  | 
  \begin{description}
 | 
|
509  | 
||
510  | 
  \item @{method (HOL) pat_completeness} is a specialized method to
 | 
|
511  | 
solve goals regarding the completeness of pattern matching, as  | 
|
512  | 
  required by the @{command (HOL) "function"} package (cf.\
 | 
|
513  | 
  \cite{isabelle-function}).
 | 
|
514  | 
||
515  | 
  \item @{method (HOL) relation}~@{text R} introduces a termination
 | 
|
516  | 
  proof using the relation @{text R}.  The resulting proof state will
 | 
|
517  | 
  contain goals expressing that @{text R} is wellfounded, and that the
 | 
|
518  | 
  arguments of recursive calls decrease with respect to @{text R}.
 | 
|
519  | 
Usually, this method is used as the initial proof step of manual  | 
|
520  | 
termination proofs.  | 
|
521  | 
||
522  | 
  \item @{method (HOL) "lexicographic_order"} attempts a fully
 | 
|
523  | 
automated termination proof by searching for a lexicographic  | 
|
524  | 
combination of size measures on the arguments of the function. The  | 
|
525  | 
  method accepts the same arguments as the @{method auto} method,
 | 
|
| 42930 | 526  | 
  which it uses internally to prove local descents.  The @{syntax
 | 
527  | 
  clasimpmod} modifiers are accepted (as for @{method auto}).
 | 
|
| 42908 | 528  | 
|
529  | 
In case of failure, extensive information is printed, which can help  | 
|
530  | 
  to analyse the situation (cf.\ \cite{isabelle-function}).
 | 
|
531  | 
||
532  | 
  \item @{method (HOL) "size_change"} also works on termination goals,
 | 
|
533  | 
using a variation of the size-change principle, together with a  | 
|
534  | 
  graph decomposition technique (see \cite{krauss_phd} for details).
 | 
|
535  | 
  Three kinds of orders are used internally: @{text max}, @{text min},
 | 
|
536  | 
  and @{text ms} (multiset), which is only available when the theory
 | 
|
537  | 
  @{text Multiset} is loaded. When no order kinds are given, they are
 | 
|
538  | 
tried in order. The search for a termination proof uses SAT solving  | 
|
539  | 
internally.  | 
|
540  | 
||
| 42930 | 541  | 
  For local descent proofs, the @{syntax clasimpmod} modifiers are
 | 
542  | 
  accepted (as for @{method auto}).
 | 
|
| 42908 | 543  | 
|
544  | 
  \end{description}
 | 
|
545  | 
*}  | 
|
546  | 
||
547  | 
||
548  | 
subsection {* Functions with explicit partiality *}
 | 
|
549  | 
||
550  | 
text {*
 | 
|
551  | 
  \begin{matharray}{rcl}
 | 
|
552  | 
    @{command_def (HOL) "partial_function"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
|
553  | 
    @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\
 | 
|
554  | 
  \end{matharray}
 | 
|
555  | 
||
556  | 
  @{rail "
 | 
|
557  | 
    @@{command (HOL) partial_function} @{syntax target}?
 | 
|
558  | 
      '(' @{syntax nameref} ')' @{syntax \"fixes\"} \\
 | 
|
559  | 
      @'where' @{syntax thmdecl}? @{syntax prop}
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
560  | 
"}  | 
| 26849 | 561  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
562  | 
  \begin{description}
 | 
| 42123 | 563  | 
|
| 42908 | 564  | 
  \item @{command (HOL) "partial_function"}~@{text "(mode)"} defines
 | 
565  | 
recursive functions based on fixpoints in complete partial  | 
|
566  | 
orders. No termination proof is required from the user or  | 
|
567  | 
constructed internally. Instead, the possibility of non-termination  | 
|
568  | 
is modelled explicitly in the result type, which contains an  | 
|
569  | 
explicit bottom element.  | 
|
570  | 
||
571  | 
Pattern matching and mutual recursion are currently not supported.  | 
|
572  | 
Thus, the specification consists of a single function described by a  | 
|
573  | 
single recursive equation.  | 
|
574  | 
||
575  | 
There are no fixed syntactic restrictions on the body of the  | 
|
576  | 
function, but the induced functional must be provably monotonic  | 
|
577  | 
wrt.\ the underlying order. The monotonicitity proof is performed  | 
|
578  | 
internally, and the definition is rejected when it fails. The proof  | 
|
579  | 
can be influenced by declaring hints using the  | 
|
580  | 
  @{attribute (HOL) partial_function_mono} attribute.
 | 
|
581  | 
||
582  | 
  The mandatory @{text mode} argument specifies the mode of operation
 | 
|
583  | 
of the command, which directly corresponds to a complete partial  | 
|
584  | 
order on the result type. By default, the following modes are  | 
|
585  | 
defined:  | 
|
| 26849 | 586  | 
|
| 42908 | 587  | 
  \begin{description}
 | 
588  | 
  \item @{text option} defines functions that map into the @{type
 | 
|
589  | 
  option} type. Here, the value @{term None} is used to model a
 | 
|
590  | 
  non-terminating computation. Monotonicity requires that if @{term
 | 
|
591  | 
None} is returned by a recursive call, then the overall result  | 
|
592  | 
  must also be @{term None}. This is best achieved through the use of
 | 
|
593  | 
  the monadic operator @{const "Option.bind"}.
 | 
|
594  | 
||
595  | 
  \item @{text tailrec} defines functions with an arbitrary result
 | 
|
596  | 
  type and uses the slightly degenerated partial order where @{term
 | 
|
597  | 
"undefined"} is the bottom element. Now, monotonicity requires that  | 
|
598  | 
  if @{term undefined} is returned by a recursive call, then the
 | 
|
599  | 
  overall result must also be @{term undefined}. In practice, this is
 | 
|
600  | 
only satisfied when each recursive call is a tail call, whose result  | 
|
601  | 
is directly returned. Thus, this mode of operation allows the  | 
|
602  | 
definition of arbitrary tail-recursive functions.  | 
|
603  | 
  \end{description}
 | 
|
604  | 
||
605  | 
Experienced users may define new modes by instantiating the locale  | 
|
606  | 
  @{const "partial_function_definitions"} appropriately.
 | 
|
607  | 
||
608  | 
  \item @{attribute (HOL) partial_function_mono} declares rules for
 | 
|
609  | 
use in the internal monononicity proofs of partial function  | 
|
610  | 
definitions.  | 
|
| 26849 | 611  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
612  | 
  \end{description}
 | 
| 42908 | 613  | 
|
614  | 
*}  | 
|
615  | 
||
616  | 
||
617  | 
subsection {* Old-style recursive function definitions (TFL) *}
 | 
|
618  | 
||
619  | 
text {*
 | 
|
620  | 
  The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
 | 
|
621  | 
  "recdef_tc"} for defining recursive are mostly obsolete; @{command
 | 
|
622  | 
  (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
 | 
|
623  | 
||
624  | 
  \begin{matharray}{rcl}
 | 
|
625  | 
    @{command_def (HOL) "recdef"} & : & @{text "theory \<rightarrow> theory)"} \\
 | 
|
626  | 
    @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
 | 
|
627  | 
  \end{matharray}
 | 
|
628  | 
||
629  | 
  @{rail "
 | 
|
630  | 
    @@{command (HOL) recdef} ('(' @'permissive' ')')? \\
 | 
|
631  | 
      @{syntax name} @{syntax term} (@{syntax prop} +) hints?
 | 
|
632  | 
;  | 
|
633  | 
    recdeftc @{syntax thmdecl}? tc
 | 
|
634  | 
;  | 
|
635  | 
    hints: '(' @'hints' ( recdefmod * ) ')'
 | 
|
636  | 
;  | 
|
637  | 
    recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf')
 | 
|
638  | 
      (() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod}
 | 
|
639  | 
;  | 
|
640  | 
    tc: @{syntax nameref} ('(' @{syntax nat} ')')?
 | 
|
641  | 
"}  | 
|
642  | 
||
643  | 
  \begin{description}
 | 
|
644  | 
||
645  | 
  \item @{command (HOL) "recdef"} defines general well-founded
 | 
|
646  | 
recursive functions (using the TFL package), see also  | 
|
647  | 
  \cite{isabelle-HOL}.  The ``@{text "(permissive)"}'' option tells
 | 
|
648  | 
TFL to recover from failed proof attempts, returning unfinished  | 
|
649  | 
  results.  The @{text recdef_simp}, @{text recdef_cong}, and @{text
 | 
|
650  | 
recdef_wf} hints refer to auxiliary rules to be used in the internal  | 
|
651  | 
  automated proof process of TFL.  Additional @{syntax clasimpmod}
 | 
|
| 42930 | 652  | 
declarations may be given to tune the context of the Simplifier  | 
653  | 
  (cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\
 | 
|
654  | 
  \secref{sec:classical}).
 | 
|
| 42908 | 655  | 
|
656  | 
  \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the
 | 
|
657  | 
  proof for leftover termination condition number @{text i} (default
 | 
|
658  | 
  1) as generated by a @{command (HOL) "recdef"} definition of
 | 
|
659  | 
  constant @{text c}.
 | 
|
660  | 
||
661  | 
  Note that in most cases, @{command (HOL) "recdef"} is able to finish
 | 
|
662  | 
its internal proofs without manual intervention.  | 
|
663  | 
||
664  | 
  \end{description}
 | 
|
665  | 
||
666  | 
  \medskip Hints for @{command (HOL) "recdef"} may be also declared
 | 
|
667  | 
globally, using the following attributes.  | 
|
668  | 
||
669  | 
  \begin{matharray}{rcl}
 | 
|
670  | 
    @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\
 | 
|
671  | 
    @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\
 | 
|
672  | 
    @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\
 | 
|
673  | 
  \end{matharray}
 | 
|
674  | 
||
675  | 
  @{rail "
 | 
|
676  | 
    (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} |
 | 
|
677  | 
      @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del')
 | 
|
678  | 
"}  | 
|
679  | 
*}  | 
|
680  | 
||
681  | 
||
682  | 
section {* Datatypes \label{sec:hol-datatype} *}
 | 
|
683  | 
||
684  | 
text {*
 | 
|
685  | 
  \begin{matharray}{rcl}
 | 
|
686  | 
    @{command_def (HOL) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
687  | 
    @{command_def (HOL) "rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
 | 
|
688  | 
  \end{matharray}
 | 
|
689  | 
||
690  | 
  @{rail "
 | 
|
691  | 
    @@{command (HOL) datatype} (spec + @'and')
 | 
|
692  | 
;  | 
|
693  | 
    @@{command (HOL) rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +)
 | 
|
694  | 
;  | 
|
695  | 
||
696  | 
    spec: @{syntax parname}? @{syntax typespec} @{syntax mixfix}? '=' (cons + '|')
 | 
|
697  | 
;  | 
|
698  | 
    cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}?
 | 
|
699  | 
"}  | 
|
700  | 
||
701  | 
  \begin{description}
 | 
|
702  | 
||
703  | 
  \item @{command (HOL) "datatype"} defines inductive datatypes in
 | 
|
704  | 
HOL.  | 
|
705  | 
||
706  | 
  \item @{command (HOL) "rep_datatype"} represents existing types as
 | 
|
| 42909 | 707  | 
datatypes.  | 
708  | 
||
709  | 
  For foundational reasons, some basic types such as @{typ nat}, @{typ
 | 
|
710  | 
  "'a \<times> 'b"}, @{typ "'a + 'b"}, @{typ bool} and @{typ unit} are
 | 
|
711  | 
  introduced by more primitive means using @{command_ref typedef}.  To
 | 
|
712  | 
  recover the rich infrastructure of @{command datatype} (e.g.\ rules
 | 
|
713  | 
  for @{method cases} and @{method induct} and the primitive recursion
 | 
|
714  | 
combinators), such types may be represented as actual datatypes  | 
|
715  | 
later. This is done by specifying the constructors of the desired  | 
|
716  | 
type, and giving a proof of the induction rule, distinctness and  | 
|
717  | 
injectivity of constructors.  | 
|
718  | 
||
719  | 
  For example, see @{file "~~/src/HOL/Sum_Type.thy"} for the
 | 
|
720  | 
representation of the primitive sum type as fully-featured datatype.  | 
|
| 42908 | 721  | 
|
722  | 
  \end{description}
 | 
|
723  | 
||
| 42909 | 724  | 
  The generated rules for @{method induct} and @{method cases} provide
 | 
725  | 
case names according to the given constructors, while parameters are  | 
|
726  | 
  named after the types (see also \secref{sec:cases-induct}).
 | 
|
| 42908 | 727  | 
|
728  | 
  See \cite{isabelle-HOL} for more details on datatypes, but beware of
 | 
|
729  | 
the old-style theory syntax being used there! Apart from proper  | 
|
730  | 
proof methods for case-analysis and induction, there are also  | 
|
731  | 
  emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
 | 
|
732  | 
  induct_tac} available, see \secref{sec:hol-induct-tac}; these admit
 | 
|
733  | 
to refer directly to the internal structure of subgoals (including  | 
|
734  | 
internally bound parameters).  | 
|
| 26849 | 735  | 
*}  | 
736  | 
||
737  | 
||
| 
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
 | 
738  | 
subsubsection {* 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
 | 
739  | 
|
| 
 
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
 | 
740  | 
text {* We define a type of finite sequences, with slightly different
 | 
| 
 
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
 | 
741  | 
  names than the existing @{typ "'a list"} that is already in @{theory
 | 
| 
 
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
 | 
742  | 
Main}: *}  | 
| 
 
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
 | 
743  | 
|
| 
 
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
 | 
744  | 
datatype 'a seq = Empty | Seq 'a "'a seq"  | 
| 
 
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
 | 
745  | 
|
| 
 
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
 | 
746  | 
text {* 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
 | 
747  | 
|
| 
 
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
 | 
748  | 
lemma "Seq x xs \<noteq> xs"  | 
| 
 
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
 | 
749  | 
proof (induct xs arbitrary: x)  | 
| 
 
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
 | 
750  | 
case 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
 | 
751  | 
  txt {* 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
 | 
752  | 
    properties of the datatype are already declared as @{attribute
 | 
| 
 
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
 | 
753  | 
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
 | 
754  | 
show "Seq x Empty \<noteq> 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
 | 
755  | 
by simp  | 
| 
 
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
 | 
756  | 
next  | 
| 
 
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
 | 
757  | 
case (Seq y ys)  | 
| 
 
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
 | 
758  | 
  txt {* 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
 | 
759  | 
show "Seq x (Seq y ys) \<noteq> Seq y ys"  | 
| 
 
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
 | 
760  | 
using `Seq y ys \<noteq> ys` by simp  | 
| 
 
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
 | 
761  | 
qed  | 
| 
 
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
 | 
762  | 
|
| 
 
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
 | 
763  | 
text {* 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
 | 
764  | 
|
| 
 
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
 | 
765  | 
lemma "Seq x xs \<noteq> xs"  | 
| 
 
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
 | 
766  | 
by (induct xs arbitrary: x) simp_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
 | 
767  | 
|
| 
 
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
 | 
768  | 
|
| 26849 | 769  | 
section {* Records \label{sec:hol-record} *}
 | 
770  | 
||
771  | 
text {*
 | 
|
772  | 
In principle, records merely generalize the concept of tuples, where  | 
|
773  | 
components may be addressed by labels instead of just position. The  | 
|
774  | 
logical infrastructure of records in Isabelle/HOL is slightly more  | 
|
775  | 
advanced, though, supporting truly extensible record schemes. This  | 
|
776  | 
admits operations that are polymorphic with respect to record  | 
|
777  | 
extension, yielding ``object-oriented'' effects like (single)  | 
|
778  | 
  inheritance.  See also \cite{NaraschewskiW-TPHOLs98} for more
 | 
|
779  | 
details on object-oriented verification and record subtyping in HOL.  | 
|
780  | 
*}  | 
|
781  | 
||
782  | 
||
783  | 
subsection {* Basic concepts *}
 | 
|
784  | 
||
785  | 
text {*
 | 
|
786  | 
  Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
 | 
|
787  | 
at the level of terms and types. The notation is as follows:  | 
|
788  | 
||
789  | 
  \begin{center}
 | 
|
790  | 
  \begin{tabular}{l|l|l}
 | 
|
791  | 
& record terms & record types \\ \hline  | 
|
792  | 
    fixed & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\
 | 
|
793  | 
    schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
 | 
|
794  | 
      @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
 | 
|
795  | 
  \end{tabular}
 | 
|
796  | 
  \end{center}
 | 
|
797  | 
||
798  | 
  \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
 | 
|
799  | 
"(| x = a |)"}.  | 
|
800  | 
||
801  | 
  A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value
 | 
|
802  | 
  @{text a} and field @{text y} of value @{text b}.  The corresponding
 | 
|
803  | 
  type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
 | 
|
804  | 
  and @{text "b :: B"}.
 | 
|
805  | 
||
806  | 
  A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields
 | 
|
807  | 
  @{text x} and @{text y} as before, but also possibly further fields
 | 
|
808  | 
  as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
 | 
|
809  | 
  of the syntax).  The improper field ``@{text "\<dots>"}'' of a record
 | 
|
810  | 
  scheme is called the \emph{more part}.  Logically it is just a free
 | 
|
811  | 
variable, which is occasionally referred to as ``row variable'' in  | 
|
812  | 
the literature. The more part of a record scheme may be  | 
|
813  | 
instantiated by zero or more further components. For example, the  | 
|
814  | 
  previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z =
 | 
|
| 26852 | 815  | 
  c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
 | 
| 26849 | 816  | 
Fixed records are special instances of record schemes, where  | 
817  | 
  ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
 | 
|
818  | 
  element.  In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
 | 
|
819  | 
  for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
 | 
|
| 42123 | 820  | 
|
| 26849 | 821  | 
\medskip Two key observations make extensible records in a simply  | 
822  | 
typed language like HOL work out:  | 
|
823  | 
||
824  | 
  \begin{enumerate}
 | 
|
825  | 
||
826  | 
\item the more part is internalized, as a free term or type  | 
|
827  | 
variable,  | 
|
828  | 
||
| 26852 | 829  | 
\item field names are externalized, they cannot be accessed within  | 
830  | 
the logic as first-class values.  | 
|
| 26849 | 831  | 
|
832  | 
  \end{enumerate}
 | 
|
833  | 
||
834  | 
\medskip In Isabelle/HOL record types have to be defined explicitly,  | 
|
835  | 
fixing their field names and types, and their (optional) parent  | 
|
836  | 
record. Afterwards, records may be formed using above syntax, while  | 
|
837  | 
obeying the canonical order of fields as given by their declaration.  | 
|
838  | 
The record package provides several standard operations like  | 
|
839  | 
selectors and updates. The common setup for various generic proof  | 
|
840  | 
tools enable succinct reasoning patterns. See also the Isabelle/HOL  | 
|
841  | 
  tutorial \cite{isabelle-hol-book} for further instructions on using
 | 
|
842  | 
records in practice.  | 
|
843  | 
*}  | 
|
844  | 
||
845  | 
||
846  | 
subsection {* Record specifications *}
 | 
|
847  | 
||
848  | 
text {*
 | 
|
849  | 
  \begin{matharray}{rcl}
 | 
|
| 
28761
 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 
wenzelm 
parents: 
28760 
diff
changeset
 | 
850  | 
    @{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
| 26849 | 851  | 
  \end{matharray}
 | 
852  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
853  | 
  @{rail "
 | 
| 42705 | 854  | 
    @@{command (HOL) record} @{syntax typespec_sorts} '=' \\
 | 
| 42704 | 855  | 
      (@{syntax type} '+')? (@{syntax constdecl} +)
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
856  | 
"}  | 
| 26849 | 857  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
858  | 
  \begin{description}
 | 
| 26849 | 859  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
860  | 
  \item @{command (HOL) "record"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1
 | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
861  | 
  \<dots> c\<^sub>n :: \<sigma>\<^sub>n"} defines extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"},
 | 
| 26849 | 862  | 
  derived from the optional parent record @{text "\<tau>"} by adding new
 | 
863  | 
  field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
 | 
|
864  | 
||
865  | 
  The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
 | 
|
866  | 
  covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
 | 
|
867  | 
  \<alpha>\<^sub>m"}.  Type constructor @{text t} has to be new, while @{text
 | 
|
868  | 
\<tau>} needs to specify an instance of an existing record type. At  | 
|
869  | 
  least one new field @{text "c\<^sub>i"} has to be specified.
 | 
|
870  | 
Basically, field names need to belong to a unique record. This is  | 
|
871  | 
not a real restriction in practice, since fields are qualified by  | 
|
872  | 
the record name internally.  | 
|
873  | 
||
874  | 
  The parent record specification @{text \<tau>} is optional; if omitted
 | 
|
875  | 
  @{text t} becomes a root record.  The hierarchy of all records
 | 
|
876  | 
declared within a theory context forms a forest structure, i.e.\ a  | 
|
877  | 
set of trees starting with a root record each. There is no way to  | 
|
878  | 
merge multiple parent records!  | 
|
879  | 
||
880  | 
  For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
 | 
|
881  | 
  type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
 | 
|
882  | 
  \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
 | 
|
883  | 
"(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for  | 
|
884  | 
  @{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> ::
 | 
|
885  | 
\<zeta>\<rparr>"}.  | 
|
886  | 
||
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
887  | 
  \end{description}
 | 
| 26849 | 888  | 
*}  | 
889  | 
||
890  | 
||
891  | 
subsection {* Record operations *}
 | 
|
892  | 
||
893  | 
text {*
 | 
|
894  | 
Any record definition of the form presented above produces certain  | 
|
895  | 
standard operations. Selectors and updates are provided for any  | 
|
896  | 
  field, including the improper one ``@{text more}''.  There are also
 | 
|
897  | 
cumulative record constructor functions. To simplify the  | 
|
898  | 
  presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
 | 
|
899  | 
  \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
 | 
|
900  | 
\<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.  | 
|
901  | 
||
902  | 
  \medskip \textbf{Selectors} and \textbf{updates} are available for
 | 
|
903  | 
  any field (including ``@{text more}''):
 | 
|
904  | 
||
905  | 
  \begin{matharray}{lll}
 | 
|
| 26852 | 906  | 
    @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
 | 
907  | 
    @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
 | 
|
| 26849 | 908  | 
  \end{matharray}
 | 
909  | 
||
910  | 
  There is special syntax for application of updates: @{text "r\<lparr>x :=
 | 
|
911  | 
  a\<rparr>"} abbreviates term @{text "x_update a r"}.  Further notation for
 | 
|
912  | 
  repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z :=
 | 
|
913  | 
  c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}.  Note that
 | 
|
914  | 
because of postfix notation the order of fields shown here is  | 
|
915  | 
reverse than in the actual term. Since repeated updates are just  | 
|
916  | 
  function applications, fields may be freely permuted in @{text "\<lparr>x
 | 
|
917  | 
:= a, y := b, z := c\<rparr>"}, as far as logical equality is concerned.  | 
|
918  | 
Thus commutativity of independent updates can be proven within the  | 
|
919  | 
logic for any two fields, but not as a general theorem.  | 
|
920  | 
||
921  | 
  \medskip The \textbf{make} operation provides a cumulative record
 | 
|
922  | 
constructor function:  | 
|
923  | 
||
924  | 
  \begin{matharray}{lll}
 | 
|
| 26852 | 925  | 
    @{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
 | 
| 26849 | 926  | 
  \end{matharray}
 | 
927  | 
||
928  | 
\medskip We now reconsider the case of non-root records, which are  | 
|
929  | 
derived of some parent. In general, the latter may depend on  | 
|
930  | 
  another parent as well, resulting in a list of \emph{ancestor
 | 
|
931  | 
records}. Appending the lists of fields of all ancestors results in  | 
|
932  | 
a certain field prefix. The record package automatically takes care  | 
|
933  | 
of this by lifting operations over this context of ancestor fields.  | 
|
934  | 
  Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
 | 
|
935  | 
  fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"},
 | 
|
936  | 
the above record operations will get the following types:  | 
|
937  | 
||
| 26852 | 938  | 
\medskip  | 
939  | 
  \begin{tabular}{lll}
 | 
|
940  | 
    @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
 | 
|
| 42123 | 941  | 
    @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow>
 | 
| 26852 | 942  | 
\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>  | 
943  | 
\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\  | 
|
944  | 
    @{text "t.make"} & @{text "::"} & @{text "\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow>
 | 
|
945  | 
\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\  | 
|
946  | 
  \end{tabular}
 | 
|
947  | 
\medskip  | 
|
| 26849 | 948  | 
|
| 26852 | 949  | 
\noindent Some further operations address the extension aspect of a  | 
| 26849 | 950  | 
  derived record scheme specifically: @{text "t.fields"} produces a
 | 
951  | 
record fragment consisting of exactly the new fields introduced here  | 
|
952  | 
  (the result may serve as a more part elsewhere); @{text "t.extend"}
 | 
|
953  | 
  takes a fixed record and adds a given more part; @{text
 | 
|
954  | 
"t.truncate"} restricts a record scheme to a fixed record.  | 
|
955  | 
||
| 26852 | 956  | 
\medskip  | 
957  | 
  \begin{tabular}{lll}
 | 
|
958  | 
    @{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
 | 
|
959  | 
    @{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
 | 
|
960  | 
\<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\  | 
|
961  | 
    @{text "t.truncate"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
 | 
|
962  | 
  \end{tabular}
 | 
|
963  | 
\medskip  | 
|
| 26849 | 964  | 
|
965  | 
  \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
 | 
|
966  | 
for root records.  | 
|
967  | 
*}  | 
|
968  | 
||
969  | 
||
970  | 
subsection {* Derived rules and proof tools *}
 | 
|
971  | 
||
972  | 
text {*
 | 
|
973  | 
The record package proves several results internally, declaring  | 
|
974  | 
these facts to appropriate proof tools. This enables users to  | 
|
975  | 
reason about record structures quite conveniently. Assume that  | 
|
976  | 
  @{text t} is a record type as specified above.
 | 
|
977  | 
||
978  | 
  \begin{enumerate}
 | 
|
| 42123 | 979  | 
|
| 26849 | 980  | 
\item Standard conversions for selectors or updates applied to  | 
981  | 
record constructor terms are made part of the default Simplifier  | 
|
982  | 
context; thus proofs by reduction of basic operations merely require  | 
|
983  | 
  the @{method simp} method without further arguments.  These rules
 | 
|
984  | 
  are available as @{text "t.simps"}, too.
 | 
|
| 42123 | 985  | 
|
| 26849 | 986  | 
\item Selectors applied to updated records are automatically reduced  | 
987  | 
by an internal simplification procedure, which is also part of the  | 
|
988  | 
standard Simplifier setup.  | 
|
989  | 
||
990  | 
  \item Inject equations of a form analogous to @{prop "(x, y) = (x',
 | 
|
991  | 
y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical  | 
|
992  | 
  Reasoner as @{attribute iff} rules.  These rules are available as
 | 
|
993  | 
  @{text "t.iffs"}.
 | 
|
994  | 
||
995  | 
  \item The introduction rule for record equality analogous to @{text
 | 
|
996  | 
"x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier,  | 
|
997  | 
  and as the basic rule context as ``@{attribute intro}@{text "?"}''.
 | 
|
998  | 
  The rule is called @{text "t.equality"}.
 | 
|
999  | 
||
1000  | 
\item Representations of arbitrary record expressions as canonical  | 
|
1001  | 
  constructor terms are provided both in @{method cases} and @{method
 | 
|
1002  | 
induct} format (cf.\ the generic proof methods of the same name,  | 
|
1003  | 
  \secref{sec:cases-induct}).  Several variations are available, for
 | 
|
1004  | 
fixed records, record schemes, more parts etc.  | 
|
| 42123 | 1005  | 
|
| 26849 | 1006  | 
The generic proof methods are sufficiently smart to pick the most  | 
1007  | 
sensible rule according to the type of the indicated record  | 
|
1008  | 
  expression: users just need to apply something like ``@{text "(cases
 | 
|
1009  | 
r)"}'' to a certain proof problem.  | 
|
1010  | 
||
1011  | 
  \item The derived record operations @{text "t.make"}, @{text
 | 
|
1012  | 
  "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
 | 
|
1013  | 
treated automatically, but usually need to be expanded by hand,  | 
|
1014  | 
  using the collective fact @{text "t.defs"}.
 | 
|
1015  | 
||
1016  | 
  \end{enumerate}
 | 
|
1017  | 
*}  | 
|
1018  | 
||
1019  | 
||
| 42911 | 1020  | 
subsubsection {* Examples *}
 | 
1021  | 
||
1022  | 
text {* See @{file "~~/src/HOL/ex/Records.thy"}, for example. *}
 | 
|
1023  | 
||
1024  | 
||
| 42908 | 1025  | 
section {* Adhoc tuples *}
 | 
| 26849 | 1026  | 
|
1027  | 
text {*
 | 
|
1028  | 
  \begin{matharray}{rcl}
 | 
|
| 42908 | 1029  | 
    @{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\
 | 
| 26849 | 1030  | 
  \end{matharray}
 | 
1031  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1032  | 
  @{rail "
 | 
| 42908 | 1033  | 
    @@{attribute (HOL) split_format} ('(' 'complete' ')')?
 | 
1034  | 
"}  | 
|
1035  | 
||
1036  | 
  \begin{description}
 | 
|
1037  | 
||
1038  | 
  \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes
 | 
|
1039  | 
arguments in function applications to be represented canonically  | 
|
1040  | 
according to their tuple type structure.  | 
|
1041  | 
||
1042  | 
Note that this operation tends to invent funny names for new local  | 
|
1043  | 
parameters introduced.  | 
|
1044  | 
||
1045  | 
  \end{description}
 | 
|
1046  | 
*}  | 
|
1047  | 
||
1048  | 
||
1049  | 
section {* Typedef axiomatization \label{sec:hol-typedef} *}
 | 
|
1050  | 
||
1051  | 
text {* A Gordon/HOL-style type definition is a certain axiom scheme
 | 
|
1052  | 
that identifies a new type with a subset of an existing type. More  | 
|
1053  | 
precisely, the new type is defined by exhibiting an existing type  | 
|
1054  | 
  @{text \<tau>}, a set @{text "A :: \<tau> set"}, and a theorem that proves
 | 
|
1055  | 
  @{prop "\<exists>x. x \<in> A"}.  Thus @{text A} is a non-empty subset of @{text
 | 
|
1056  | 
\<tau>}, and the new type denotes this subset. New functions are  | 
|
1057  | 
postulated that establish an isomorphism between the new type and  | 
|
1058  | 
  the subset.  In general, the type @{text \<tau>} may involve type
 | 
|
1059  | 
  variables @{text "\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n"} which means that the type definition
 | 
|
1060  | 
  produces a type constructor @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} depending on
 | 
|
1061  | 
those type arguments.  | 
|
1062  | 
||
1063  | 
The axiomatization can be considered a ``definition'' in the sense  | 
|
1064  | 
of the particular set-theoretic interpretation of HOL  | 
|
1065  | 
  \cite{pitts93}, where the universe of types is required to be
 | 
|
1066  | 
downwards-closed wrt.\ arbitrary non-empty subsets. Thus genuinely  | 
|
1067  | 
  new types introduced by @{command "typedef"} stay within the range
 | 
|
1068  | 
  of HOL models by construction.  Note that @{command_ref
 | 
|
1069  | 
type_synonym} from Isabelle/Pure merely introduces syntactic  | 
|
1070  | 
abbreviations, without any logical significance.  | 
|
1071  | 
||
1072  | 
  \begin{matharray}{rcl}
 | 
|
1073  | 
    @{command_def (HOL) "typedef"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
 | 
|
1074  | 
  \end{matharray}
 | 
|
1075  | 
||
1076  | 
  @{rail "
 | 
|
1077  | 
    @@{command (HOL) typedef} alt_name? abs_type '=' rep_set
 | 
|
| 26849 | 1078  | 
;  | 
1079  | 
||
| 42908 | 1080  | 
    alt_name: '(' (@{syntax name} | @'open' | @'open' @{syntax name}) ')'
 | 
| 26849 | 1081  | 
;  | 
| 42908 | 1082  | 
    abs_type: @{syntax typespec_sorts} @{syntax mixfix}?
 | 
1083  | 
;  | 
|
1084  | 
    rep_set: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})?
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1085  | 
"}  | 
| 26849 | 1086  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1087  | 
  \begin{description}
 | 
| 26849 | 1088  | 
|
| 42908 | 1089  | 
  \item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"}
 | 
1090  | 
axiomatizes a type definition in the background theory of the  | 
|
1091  | 
current context, depending on a non-emptiness result of the set  | 
|
1092  | 
  @{text A} that needs to be proven here.  The set @{text A} may
 | 
|
1093  | 
  contain type variables @{text "\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n"} as specified on the LHS,
 | 
|
1094  | 
but no term variables.  | 
|
1095  | 
||
1096  | 
Even though a local theory specification, the newly introduced type  | 
|
1097  | 
constructor cannot depend on parameters or assumptions of the  | 
|
1098  | 
context: this is structurally impossible in HOL. In contrast, the  | 
|
1099  | 
non-emptiness proof may use local assumptions in unusual situations,  | 
|
1100  | 
which could result in different interpretations in target contexts:  | 
|
1101  | 
  the meaning of the bijection between the representing set @{text A}
 | 
|
1102  | 
  and the new type @{text t} may then change in different application
 | 
|
1103  | 
contexts.  | 
|
1104  | 
||
1105  | 
  By default, @{command (HOL) "typedef"} defines both a type
 | 
|
1106  | 
  constructor @{text t} for the new type, and a term constant @{text
 | 
|
1107  | 
  t} for the representing set within the old type.  Use the ``@{text
 | 
|
1108  | 
"(open)"}'' option to suppress a separate constant definition  | 
|
1109  | 
  altogether.  The injection from type to set is called @{text Rep_t},
 | 
|
1110  | 
  its inverse @{text Abs_t}, unless explicit @{keyword (HOL)
 | 
|
1111  | 
"morphisms"} specification provides alternative names.  | 
|
| 26849 | 1112  | 
|
| 42908 | 1113  | 
  The core axiomatization uses the locale predicate @{const
 | 
1114  | 
type_definition} as defined in Isabelle/HOL. Various basic  | 
|
1115  | 
consequences of that are instantiated accordingly, re-using the  | 
|
1116  | 
locale facts with names derived from the new type constructor. Thus  | 
|
1117  | 
  the generic @{thm type_definition.Rep} is turned into the specific
 | 
|
1118  | 
  @{text "Rep_t"}, for example.
 | 
|
1119  | 
||
1120  | 
  Theorems @{thm type_definition.Rep}, @{thm
 | 
|
1121  | 
  type_definition.Rep_inverse}, and @{thm type_definition.Abs_inverse}
 | 
|
1122  | 
provide the most basic characterization as a corresponding  | 
|
1123  | 
injection/surjection pair (in both directions). The derived rules  | 
|
1124  | 
  @{thm type_definition.Rep_inject} and @{thm
 | 
|
1125  | 
type_definition.Abs_inject} provide a more convenient version of  | 
|
1126  | 
injectivity, suitable for automated proof tools (e.g.\ in  | 
|
1127  | 
  declarations involving @{attribute simp} or @{attribute iff}).
 | 
|
1128  | 
  Furthermore, the rules @{thm type_definition.Rep_cases}~/ @{thm
 | 
|
1129  | 
  type_definition.Rep_induct}, and @{thm type_definition.Abs_cases}~/
 | 
|
1130  | 
  @{thm type_definition.Abs_induct} provide alternative views on
 | 
|
1131  | 
surjectivity. These rules are already declared as set or type rules  | 
|
1132  | 
  for the generic @{method cases} and @{method induct} methods,
 | 
|
1133  | 
respectively.  | 
|
1134  | 
||
1135  | 
An alternative name for the set definition (and other derived  | 
|
1136  | 
entities) may be specified in parentheses; the default is to use  | 
|
1137  | 
  @{text t} directly.
 | 
|
| 26849 | 1138  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1139  | 
  \end{description}
 | 
| 26849 | 1140  | 
|
| 42908 | 1141  | 
  \begin{warn}
 | 
1142  | 
  If you introduce a new type axiomatically, i.e.\ via @{command_ref
 | 
|
1143  | 
  typedecl} and @{command_ref axiomatization}, the minimum requirement
 | 
|
1144  | 
is that it has a non-empty model, to avoid immediate collapse of the  | 
|
1145  | 
HOL logic. Moreover, one needs to demonstrate that the  | 
|
1146  | 
interpretation of such free-form axiomatizations can coexist with  | 
|
1147  | 
  that of the regular @{command_def typedef} scheme, and any extension
 | 
|
1148  | 
that other people might have introduced elsewhere (e.g.\ in HOLCF  | 
|
1149  | 
  \cite{MuellerNvOS99}).
 | 
|
1150  | 
  \end{warn}
 | 
|
1151  | 
*}  | 
|
1152  | 
||
1153  | 
subsubsection {* Examples *}
 | 
|
1154  | 
||
1155  | 
text {* Type definitions permit the introduction of abstract data
 | 
|
1156  | 
types in a safe way, namely by providing models based on already  | 
|
1157  | 
  existing types.  Given some abstract axiomatic description @{text P}
 | 
|
1158  | 
of a type, this involves two steps:  | 
|
1159  | 
||
1160  | 
  \begin{enumerate}
 | 
|
1161  | 
||
1162  | 
  \item Find an appropriate type @{text \<tau>} and subset @{text A} which
 | 
|
1163  | 
  has the desired properties @{text P}, and make a type definition
 | 
|
1164  | 
based on this representation.  | 
|
1165  | 
||
1166  | 
  \item Prove that @{text P} holds for @{text \<tau>} by lifting @{text P}
 | 
|
1167  | 
from the representation.  | 
|
| 26849 | 1168  | 
|
| 42908 | 1169  | 
  \end{enumerate}
 | 
1170  | 
||
1171  | 
You can later forget about the representation and work solely in  | 
|
1172  | 
  terms of the abstract properties @{text P}.
 | 
|
1173  | 
||
1174  | 
\medskip The following trivial example pulls a three-element type  | 
|
1175  | 
into existence within the formal logical environment of HOL. *}  | 
|
1176  | 
||
1177  | 
typedef three = "{(True, True), (True, False), (False, True)}"
 | 
|
1178  | 
by blast  | 
|
1179  | 
||
1180  | 
definition "One = Abs_three (True, True)"  | 
|
1181  | 
definition "Two = Abs_three (True, False)"  | 
|
1182  | 
definition "Three = Abs_three (False, True)"  | 
|
1183  | 
||
1184  | 
lemma three_distinct: "One \<noteq> Two" "One \<noteq> Three" "Two \<noteq> Three"  | 
|
1185  | 
by (simp_all add: One_def Two_def Three_def Abs_three_inject three_def)  | 
|
1186  | 
||
1187  | 
lemma three_cases:  | 
|
1188  | 
fixes x :: three obtains "x = One" | "x = Two" | "x = Three"  | 
|
1189  | 
by (cases x) (auto simp: One_def Two_def Three_def Abs_three_inject three_def)  | 
|
1190  | 
||
1191  | 
text {* Note that such trivial constructions are better done with
 | 
|
1192  | 
  derived specification mechanisms such as @{command datatype}: *}
 | 
|
1193  | 
||
1194  | 
datatype three' = One' | Two' | Three'  | 
|
1195  | 
||
1196  | 
text {* This avoids re-doing basic definitions and proofs from the
 | 
|
1197  | 
  primitive @{command typedef} above. *}
 | 
|
| 26849 | 1198  | 
|
1199  | 
||
| 41396 | 1200  | 
section {* Functorial structure of types *}
 | 
1201  | 
||
1202  | 
text {*
 | 
|
1203  | 
  \begin{matharray}{rcl}
 | 
|
| 
41505
 
6d19301074cf
"enriched_type" replaces less specific "type_lifting"
 
haftmann 
parents: 
41396 
diff
changeset
 | 
1204  | 
    @{command_def (HOL) "enriched_type"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
 | 
| 41396 | 1205  | 
  \end{matharray}
 | 
1206  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1207  | 
  @{rail "
 | 
| 42617 | 1208  | 
    @@{command (HOL) enriched_type} (@{syntax name} ':')? @{syntax term}
 | 
| 41396 | 1209  | 
;  | 
| 42617 | 1210  | 
"}  | 
| 41396 | 1211  | 
|
1212  | 
  \begin{description}
 | 
|
1213  | 
||
| 42617 | 1214  | 
  \item @{command (HOL) "enriched_type"}~@{text "prefix: m"} allows to
 | 
1215  | 
prove and register properties about the functorial structure of type  | 
|
1216  | 
constructors. These properties then can be used by other packages  | 
|
1217  | 
to deal with those type constructors in certain type constructions.  | 
|
1218  | 
Characteristic theorems are noted in the current local theory. By  | 
|
1219  | 
default, they are prefixed with the base name of the type  | 
|
1220  | 
constructor, an explicit prefix can be given alternatively.  | 
|
| 41396 | 1221  | 
|
1222  | 
  The given term @{text "m"} is considered as \emph{mapper} for the
 | 
|
1223  | 
corresponding type constructor and must conform to the following  | 
|
1224  | 
type pattern:  | 
|
1225  | 
||
1226  | 
  \begin{matharray}{lll}
 | 
|
1227  | 
    @{text "m"} & @{text "::"} &
 | 
|
1228  | 
      @{text "\<sigma>\<^isub>1 \<Rightarrow> \<dots> \<sigma>\<^isub>k \<Rightarrow> (\<^vec>\<alpha>\<^isub>n) t \<Rightarrow> (\<^vec>\<beta>\<^isub>n) t"} \\
 | 
|
1229  | 
  \end{matharray}
 | 
|
1230  | 
||
1231  | 
  \noindent where @{text t} is the type constructor, @{text
 | 
|
1232  | 
  "\<^vec>\<alpha>\<^isub>n"} and @{text "\<^vec>\<beta>\<^isub>n"} are distinct
 | 
|
1233  | 
  type variables free in the local theory and @{text "\<sigma>\<^isub>1"},
 | 
|
1234  | 
  \ldots, @{text "\<sigma>\<^isub>k"} is a subsequence of @{text "\<alpha>\<^isub>1 \<Rightarrow>
 | 
|
1235  | 
  \<beta>\<^isub>1"}, @{text "\<beta>\<^isub>1 \<Rightarrow> \<alpha>\<^isub>1"}, \ldots,
 | 
|
1236  | 
  @{text "\<alpha>\<^isub>n \<Rightarrow> \<beta>\<^isub>n"}, @{text "\<beta>\<^isub>n \<Rightarrow>
 | 
|
1237  | 
\<alpha>\<^isub>n"}.  | 
|
1238  | 
||
1239  | 
  \end{description}
 | 
|
1240  | 
*}  | 
|
1241  | 
||
| 
43993
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1242  | 
section {* Quotient types *}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1243  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1244  | 
text {*
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1245  | 
The quotient package defines a new quotient type given a raw type  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1246  | 
and a partial equivalence relation.  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1247  | 
It also includes automation for transporting definitions and theorems.  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1248  | 
It can automatically produce definitions and theorems on the quotient type,  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1249  | 
given the corresponding constants and facts on the raw type.  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1250  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1251  | 
  \begin{matharray}{rcl}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1252  | 
    @{command_def (HOL) "quotient_type"} & : & @{text "local_theory \<rightarrow> proof(prove)"}\\
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1253  | 
    @{command_def (HOL) "quotient_definition"} & : & @{text "local_theory \<rightarrow> proof(prove)"}\\
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1254  | 
    @{command_def (HOL) "print_quotmaps"} & : & @{text "context \<rightarrow>"}\\
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1255  | 
    @{command_def (HOL) "print_quotients"} & : & @{text "context \<rightarrow>"}\\
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1256  | 
    @{command_def (HOL) "print_quotconsts"} & : & @{text "context \<rightarrow>"}\\
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1257  | 
  \end{matharray}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1258  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1259  | 
  @{rail "
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1260  | 
    @@{command (HOL) quotient_type} (spec + @'and');
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1261  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1262  | 
    spec: @{syntax typespec} @{syntax mixfix}? '=' \\
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1263  | 
     @{syntax type} '/' ('partial' ':')? @{syntax term}; 
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1264  | 
"}  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1265  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1266  | 
  @{rail "
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1267  | 
    @@{command (HOL) quotient_definition} constdecl? @{syntax thmdecl}? \\
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1268  | 
    @{syntax term} 'is' @{syntax term};
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1269  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1270  | 
    constdecl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}?
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1271  | 
"}  | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1272  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1273  | 
  \begin{description}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1274  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1275  | 
  \item @{command (HOL) "quotient_type"} defines quotient types.
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1276  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1277  | 
  \item @{command (HOL) "quotient_definition"} defines a constant on the quotient type.
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1278  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1279  | 
  \item @{command (HOL) "print_quotmaps"} prints quotient map functions.
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1280  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1281  | 
  \item @{command (HOL) "print_quotients"} prints quotients.
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1282  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1283  | 
  \item @{command (HOL) "print_quotconsts"} prints quotient constants.
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1284  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1285  | 
  \end{description}
 | 
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1286  | 
|
| 
 
b141d7a3d4e3
rudimentary documentation of the quotient package in the isar reference manual
 
bulwahn 
parents: 
43914 
diff
changeset
 | 
1287  | 
*}  | 
| 41396 | 1288  | 
|
| 43994 | 1289  | 
section {* Coercive subtyping *}
 | 
1290  | 
||
1291  | 
text {*
 | 
|
1292  | 
  \begin{matharray}{rcl}
 | 
|
1293  | 
    @{attribute_def (HOL) coercion} & : & @{text attribute} \\
 | 
|
1294  | 
    @{attribute_def (HOL) coercion_enabled} & : & @{text attribute} \\
 | 
|
1295  | 
    @{attribute_def (HOL) coercion_map} & : & @{text attribute} \\
 | 
|
1296  | 
  \end{matharray}
 | 
|
1297  | 
||
1298  | 
  @{rail "
 | 
|
1299  | 
    @@{attribute (HOL) coercion} (@{syntax term})?
 | 
|
1300  | 
;  | 
|
1301  | 
"}  | 
|
1302  | 
  @{rail "
 | 
|
1303  | 
    @@{attribute (HOL) coercion_map} (@{syntax term})?
 | 
|
1304  | 
;  | 
|
1305  | 
"}  | 
|
1306  | 
||
1307  | 
Coercive subtyping allows the user to omit explicit type conversions,  | 
|
1308  | 
  also called \emph{coercions}.  Type inference will add them as
 | 
|
1309  | 
necessary when parsing a term. See  | 
|
1310  | 
  \cite{traytel-berghofer-nipkow-2011} for details.
 | 
|
1311  | 
||
1312  | 
  \begin{description}
 | 
|
1313  | 
||
1314  | 
  \item @{attribute (HOL) "coercion"}~@{text "f"} registers a new
 | 
|
1315  | 
  coercion function @{text "f :: \<sigma>\<^isub>1 \<Rightarrow>
 | 
|
1316  | 
  \<sigma>\<^isub>2"} where @{text "\<sigma>\<^isub>1"} and @{text
 | 
|
1317  | 
"\<sigma>\<^isub>2"} are nullary type constructors. Coercions are  | 
|
1318  | 
composed by the inference algorithm if needed. Note that the type  | 
|
1319  | 
inference algorithm is complete only if the registered coercions form  | 
|
1320  | 
a lattice.  | 
|
1321  | 
||
1322  | 
||
1323  | 
  \item @{attribute (HOL) "coercion_map"}~@{text "map"} registers a new
 | 
|
1324  | 
map function to lift coercions through type constructors. The function  | 
|
1325  | 
  @{text "map"} must conform to the following type pattern
 | 
|
1326  | 
||
1327  | 
  \begin{matharray}{lll}
 | 
|
1328  | 
    @{text "map"} & @{text "::"} &
 | 
|
1329  | 
      @{text "f\<^isub>1 \<Rightarrow> \<dots> \<Rightarrow> f\<^isub>n \<Rightarrow> (\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>n) t \<Rightarrow> (\<beta>\<^isub>1, \<dots>, \<beta>\<^isub>n) t"} \\
 | 
|
1330  | 
  \end{matharray}
 | 
|
1331  | 
||
1332  | 
  where @{text "t"} is a type constructor and @{text "f\<^isub>i"} is of
 | 
|
1333  | 
  type @{text "\<alpha>\<^isub>i \<Rightarrow> \<beta>\<^isub>i"} or
 | 
|
1334  | 
  @{text "\<beta>\<^isub>i \<Rightarrow> \<alpha>\<^isub>i"}.
 | 
|
1335  | 
Registering a map function overwrites any existing map function for  | 
|
1336  | 
this particular type constructor.  | 
|
1337  | 
||
1338  | 
||
1339  | 
  \item @{attribute (HOL) "coercion_enabled"} enables the coercion
 | 
|
1340  | 
inference algorithm.  | 
|
1341  | 
||
1342  | 
  \end{description}
 | 
|
1343  | 
||
1344  | 
*}  | 
|
1345  | 
||
| 26849 | 1346  | 
section {* Arithmetic proof support *}
 | 
1347  | 
||
1348  | 
text {*
 | 
|
1349  | 
  \begin{matharray}{rcl}
 | 
|
| 
28761
 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 
wenzelm 
parents: 
28760 
diff
changeset
 | 
1350  | 
    @{method_def (HOL) arith} & : & @{text method} \\
 | 
| 30863 | 1351  | 
    @{attribute_def (HOL) arith} & : & @{text attribute} \\
 | 
| 
28761
 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 
wenzelm 
parents: 
28760 
diff
changeset
 | 
1352  | 
    @{attribute_def (HOL) arith_split} & : & @{text attribute} \\
 | 
| 26849 | 1353  | 
  \end{matharray}
 | 
1354  | 
||
1355  | 
  The @{method (HOL) arith} method decides linear arithmetic problems
 | 
|
1356  | 
  (on types @{text nat}, @{text int}, @{text real}).  Any current
 | 
|
1357  | 
facts are inserted into the goal before running the procedure.  | 
|
1358  | 
||
| 30863 | 1359  | 
  The @{attribute (HOL) arith} attribute declares facts that are
 | 
1360  | 
always supplied to the arithmetic provers implicitly.  | 
|
| 26849 | 1361  | 
|
| 30863 | 1362  | 
  The @{attribute (HOL) arith_split} attribute declares case split
 | 
| 30865 | 1363  | 
  rules to be expanded before @{method (HOL) arith} is invoked.
 | 
| 30863 | 1364  | 
|
1365  | 
Note that a simpler (but faster) arithmetic prover is  | 
|
1366  | 
already invoked by the Simplifier.  | 
|
| 26849 | 1367  | 
*}  | 
1368  | 
||
1369  | 
||
| 30169 | 1370  | 
section {* Intuitionistic proof search *}
 | 
1371  | 
||
1372  | 
text {*
 | 
|
1373  | 
  \begin{matharray}{rcl}
 | 
|
| 30171 | 1374  | 
    @{method_def (HOL) iprover} & : & @{text method} \\
 | 
| 30169 | 1375  | 
  \end{matharray}
 | 
1376  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1377  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1378  | 
    @@{method (HOL) iprover} ( @{syntax rulemod} * )
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1379  | 
"}  | 
| 30169 | 1380  | 
|
| 30171 | 1381  | 
  The @{method (HOL) iprover} method performs intuitionistic proof
 | 
1382  | 
search, depending on specifically declared rules from the context,  | 
|
1383  | 
or given as explicit arguments. Chained facts are inserted into the  | 
|
| 35613 | 1384  | 
goal before commencing proof search.  | 
1385  | 
||
| 30169 | 1386  | 
  Rules need to be classified as @{attribute (Pure) intro},
 | 
1387  | 
  @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the
 | 
|
1388  | 
  ``@{text "!"}'' indicator refers to ``safe'' rules, which may be
 | 
|
1389  | 
applied aggressively (without considering back-tracking later).  | 
|
1390  | 
  Rules declared with ``@{text "?"}'' are ignored in proof search (the
 | 
|
| 42626 | 1391  | 
  single-step @{method (Pure) rule} method still observes these).  An
 | 
| 30169 | 1392  | 
explicit weight annotation may be given as well; otherwise the  | 
1393  | 
number of rule premises will be taken into account here.  | 
|
1394  | 
*}  | 
|
1395  | 
||
| 
43578
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1396  | 
section {* Model Elimination and Resolution *}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1397  | 
|
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1398  | 
text {*
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1399  | 
  \begin{matharray}{rcl}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1400  | 
    @{method_def (HOL) "meson"} & : & @{text method} \\
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1401  | 
    @{method_def (HOL) "metis"} & : & @{text method} \\
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1402  | 
  \end{matharray}
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1403  | 
|
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1404  | 
  @{rail "
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1405  | 
    @@{method (HOL) meson} @{syntax thmrefs}?
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1406  | 
;  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1407  | 
|
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1408  | 
    @@{method (HOL) metis} ( '(' ('partial_types' | 'full_types' | 'no_types'
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1409  | 
                                  | @{syntax name}) ')' )? @{syntax thmrefs}?
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1410  | 
"}  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1411  | 
|
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1412  | 
  The @{method (HOL) meson} method implements Loveland's model elimination
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1413  | 
  procedure \cite{loveland-78}. See @{file "~~/src/HOL/ex/Meson_Test.thy"} for
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1414  | 
examples.  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1415  | 
|
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1416  | 
  The @{method (HOL) metis} method combines ordered resolution and ordered
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1417  | 
paramodulation to find first-order (or mildly higher-order) proofs. The first  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1418  | 
optional argument specifies a type encoding; see the Sledgehammer manual  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1419  | 
  \cite{isabelle-sledgehammer} for details. The @{file
 | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1420  | 
"~~/src/HOL/Metis_Examples"} directory contains several small theories  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1421  | 
developed to a large extent using Metis.  | 
| 
 
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
 
blanchet 
parents: 
43270 
diff
changeset
 | 
1422  | 
*}  | 
| 30169 | 1423  | 
|
| 30171 | 1424  | 
section {* Coherent Logic *}
 | 
1425  | 
||
1426  | 
text {*
 | 
|
1427  | 
  \begin{matharray}{rcl}
 | 
|
1428  | 
    @{method_def (HOL) "coherent"} & : & @{text method} \\
 | 
|
1429  | 
  \end{matharray}
 | 
|
1430  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1431  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1432  | 
    @@{method (HOL) coherent} @{syntax thmrefs}?
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1433  | 
"}  | 
| 30171 | 1434  | 
|
1435  | 
  The @{method (HOL) coherent} method solves problems of
 | 
|
1436  | 
  \emph{Coherent Logic} \cite{Bezem-Coquand:2005}, which covers
 | 
|
1437  | 
applications in confluence theory, lattice theory and projective  | 
|
| 
40800
 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 
wenzelm 
parents: 
40711 
diff
changeset
 | 
1438  | 
  geometry.  See @{file "~~/src/HOL/ex/Coherent.thy"} for some
 | 
| 30171 | 1439  | 
examples.  | 
1440  | 
*}  | 
|
1441  | 
||
1442  | 
||
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1443  | 
section {* Proving propositions *}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1444  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1445  | 
text {*
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1446  | 
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
 | 
1447  | 
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
 | 
1448  | 
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
 | 
1449  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1450  | 
  \begin{matharray}{rcl}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1451  | 
    @{command_def (HOL) "solve_direct"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1452  | 
    @{command_def (HOL) "try"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
 | 
| 43016 | 1453  | 
    @{command_def (HOL) "try_methods"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1454  | 
    @{command_def (HOL) "sledgehammer"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1455  | 
    @{command_def (HOL) "sledgehammer_params"} & : & @{text "theory \<rightarrow> theory"}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1456  | 
  \end{matharray}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1457  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1458  | 
  @{rail "
 | 
| 43040 | 1459  | 
    @@{command (HOL) try}
 | 
1460  | 
;  | 
|
1461  | 
||
| 43016 | 1462  | 
    @@{command (HOL) try_methods} ( ( ( 'simp' | 'intro' | 'elim' | 'dest' ) ':' @{syntax thmrefs} ) + ) ?
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1463  | 
      @{syntax nat}?
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1464  | 
;  | 
| 43040 | 1465  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1466  | 
    @@{command (HOL) sledgehammer} ( '[' args ']' )? facts? @{syntax nat}?
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1467  | 
;  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1468  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1469  | 
    @@{command (HOL) sledgehammer_params} ( ( '[' args ']' ) ? )
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1470  | 
;  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1471  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1472  | 
    args: ( @{syntax name} '=' value + ',' )
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1473  | 
;  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1474  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1475  | 
    facts: '(' ( ( ( ( 'add' | 'del' ) ':' ) ? @{syntax thmrefs} ) + ) ? ')'
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1476  | 
;  | 
| 
43019
 
619f16bf2150
removed FIXME: "try_methods" doesn't take a clasimpmod, because it needs to generate Isar proof text in case of success and hence needs total control over its arguments
 
blanchet 
parents: 
43016 
diff
changeset
 | 
1477  | 
"} % FIXME check args "value"  | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1478  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1479  | 
  \begin{description}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1480  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1481  | 
  \item @{command (HOL) "solve_direct"} checks whether the current subgoals can
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1482  | 
be solved directly by an existing theorem. Duplicate lemmas can be detected  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1483  | 
in this way.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1484  | 
|
| 43016 | 1485  | 
  \item @{command (HOL) "try_methods"} attempts to prove a subgoal using a combination
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1486  | 
    of standard proof methods (@{text auto}, @{text simp}, @{text blast}, etc.).
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1487  | 
    Additional facts supplied via @{text "simp:"}, @{text "intro:"},
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1488  | 
    @{text "elim:"}, and @{text "dest:"} are passed to the appropriate proof
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1489  | 
methods.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1490  | 
|
| 
43914
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1491  | 
  \item @{command (HOL) "try"} attempts to prove or disprove a subgoal
 | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1492  | 
    using a combination of provers and disprovers (@{text "solve_direct"},
 | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1493  | 
    @{text "quickcheck"}, @{text "try_methods"}, @{text "sledgehammer"},
 | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1494  | 
    @{text "nitpick"}).
 | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1495  | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1496  | 
  \item @{command (HOL) "sledgehammer"} attempts to prove a subgoal using external
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1497  | 
automatic provers (resolution provers and SMT solvers). See the Sledgehammer  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1498  | 
    manual \cite{isabelle-sledgehammer} for details.
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1499  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1500  | 
  \item @{command (HOL) "sledgehammer_params"} changes
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1501  | 
    @{command (HOL) "sledgehammer"} configuration options persistently.
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1502  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1503  | 
  \end{description}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1504  | 
*}  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1505  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1506  | 
|
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1507  | 
section {* Checking and refuting propositions *}
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1508  | 
|
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1509  | 
text {*
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1510  | 
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
 | 
1511  | 
particular assignments and systematic counterexample search. This  | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1512  | 
is supported by the following commands.  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1513  | 
|
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1514  | 
  \begin{matharray}{rcl}
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1515  | 
    @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
| 
45409
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
1516  | 
    @{command_def (HOL) "values"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1517  | 
    @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1518  | 
    @{command_def (HOL) "refute"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1519  | 
    @{command_def (HOL) "nitpick"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1520  | 
    @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1521  | 
    @{command_def (HOL) "refute_params"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1522  | 
    @{command_def (HOL) "nitpick_params"} & : & @{text "theory \<rightarrow> theory"}
 | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1523  | 
  \end{matharray}
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1524  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1525  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1526  | 
    @@{command (HOL) value} ( '[' name ']' )? modes? @{syntax term}
 | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1527  | 
;  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1528  | 
|
| 
45409
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
1529  | 
    @@{command (HOL) values} modes? @{syntax nat}? @{syntax term}
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
1530  | 
;  | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
1531  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1532  | 
    (@@{command (HOL) quickcheck} | @@{command (HOL) refute} | @@{command (HOL) nitpick})
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1533  | 
      ( '[' args ']' )? @{syntax nat}?
 | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1534  | 
;  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1535  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1536  | 
    (@@{command (HOL) quickcheck_params} | @@{command (HOL) refute_params} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1537  | 
      @@{command (HOL) nitpick_params}) ( '[' args ']' )?
 | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1538  | 
;  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1539  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1540  | 
    modes: '(' (@{syntax name} +) ')'
 | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1541  | 
;  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1542  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1543  | 
    args: ( @{syntax name} '=' value + ',' )
 | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1544  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1545  | 
"} % FIXME check "value"  | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1546  | 
|
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1547  | 
  \begin{description}
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1548  | 
|
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1549  | 
  \item @{command (HOL) "value"}~@{text t} evaluates and prints a
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1550  | 
    term; optionally @{text modes} can be specified, which are
 | 
| 42926 | 1551  | 
    appended to the current print mode; see \secref{sec:print-modes}.
 | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1552  | 
Internally, the evaluation is performed by registered evaluators,  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1553  | 
which are invoked sequentially until a result is returned.  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1554  | 
Alternatively a specific evaluator can be selected using square  | 
| 37444 | 1555  | 
brackets; typical evaluators use the current set of code equations  | 
| 42926 | 1556  | 
    to normalize and include @{text simp} for fully symbolic
 | 
1557  | 
    evaluation using the simplifier, @{text nbe} for
 | 
|
1558  | 
    \emph{normalization by evaluation} and \emph{code} for code
 | 
|
1559  | 
generation in SML.  | 
|
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1560  | 
|
| 
45409
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
1561  | 
  \item @{command (HOL) "values"}~@{text t} enumerates a set comprehension
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
1562  | 
by evaluation and prints its values up to the given number of solutions;  | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
1563  | 
    optionally @{text modes} can be specified, which are
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
1564  | 
    appended to the current print mode; see \secref{sec:print-modes}.
 | 
| 
 
5abb0e738b00
adding some documentation about the values command to the isar reference
 
bulwahn 
parents: 
45408 
diff
changeset
 | 
1565  | 
|
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1566  | 
  \item @{command (HOL) "quickcheck"} tests the current goal for
 | 
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1567  | 
counterexamples using a series of assignments for its  | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1568  | 
free variables; by default the first subgoal is tested, an other  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1569  | 
can be selected explicitly using an optional goal index.  | 
| 
40918
 
7351c6afb348
explaining quickcheck testers in the documentation
 
bulwahn 
parents: 
40800 
diff
changeset
 | 
1570  | 
Assignments can be chosen exhausting the search space upto a given  | 
| 
43914
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1571  | 
size, or using a fixed number of random assignments in the search space,  | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1572  | 
or exploring the search space symbolically using narrowing.  | 
| 
40918
 
7351c6afb348
explaining quickcheck testers in the documentation
 
bulwahn 
parents: 
40800 
diff
changeset
 | 
1573  | 
By default, quickcheck uses exhaustive testing.  | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1574  | 
A number of configuration options are supported for  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1575  | 
    @{command (HOL) "quickcheck"}, notably:
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1576  | 
|
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1577  | 
    \begin{description}
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1578  | 
|
| 
43914
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1579  | 
    \item[@{text tester}] specifies which testing approach to apply.
 | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1580  | 
      There are three testers, @{text exhaustive},
 | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1581  | 
      @{text random}, and @{text narrowing}.
 | 
| 
40918
 
7351c6afb348
explaining quickcheck testers in the documentation
 
bulwahn 
parents: 
40800 
diff
changeset
 | 
1582  | 
An unknown configuration option is treated as an argument to tester,  | 
| 
 
7351c6afb348
explaining quickcheck testers in the documentation
 
bulwahn 
parents: 
40800 
diff
changeset
 | 
1583  | 
      making @{text "tester ="} optional.
 | 
| 
43914
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1584  | 
When multiple testers are given, these are applied in parallel.  | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1585  | 
If no tester is specified, quickcheck uses the testers that are  | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1586  | 
set active, i.e., configurations  | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1587  | 
      @{text quickcheck_exhaustive_active}, @{text quickcheck_random_active},
 | 
| 
 
64819f353c53
updating documentation about quickcheck; adding information about try
 
bulwahn 
parents: 
43578 
diff
changeset
 | 
1588  | 
      @{text quickcheck_narrowing_active} are set to true.
 | 
| 40254 | 1589  | 
    \item[@{text size}] specifies the maximum size of the search space
 | 
1590  | 
for assignment values.  | 
|
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1591  | 
|
| 
42092
 
f07b373f25d3
adding documentation about the eval option in quickcheck
 
bulwahn 
parents: 
41846 
diff
changeset
 | 
1592  | 
    \item[@{text eval}] takes a term or a list of terms and evaluates
 | 
| 
 
f07b373f25d3
adding documentation about the eval option in quickcheck
 
bulwahn 
parents: 
41846 
diff
changeset
 | 
1593  | 
these terms under the variable assignment found by quickcheck.  | 
| 42123 | 1594  | 
|
| 40254 | 1595  | 
    \item[@{text iterations}] sets how many sets of assignments are
 | 
1596  | 
generated for each particular size.  | 
|
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1597  | 
|
| 40254 | 1598  | 
    \item[@{text no_assms}] specifies whether assumptions in
 | 
1599  | 
structured proofs should be ignored.  | 
|
| 35331 | 1600  | 
|
| 40254 | 1601  | 
    \item[@{text timeout}] sets the time limit in seconds.
 | 
| 
40245
 
59f011c1877a
updating documentation on quickcheck in the Isar reference
 
bulwahn 
parents: 
40171 
diff
changeset
 | 
1602  | 
|
| 40254 | 1603  | 
    \item[@{text default_type}] sets the type(s) generally used to
 | 
1604  | 
instantiate type variables.  | 
|
| 
40245
 
59f011c1877a
updating documentation on quickcheck in the Isar reference
 
bulwahn 
parents: 
40171 
diff
changeset
 | 
1605  | 
|
| 40254 | 1606  | 
    \item[@{text report}] if set quickcheck reports how many tests
 | 
1607  | 
fulfilled the preconditions.  | 
|
| 
40245
 
59f011c1877a
updating documentation on quickcheck in the Isar reference
 
bulwahn 
parents: 
40171 
diff
changeset
 | 
1608  | 
|
| 40254 | 1609  | 
    \item[@{text quiet}] if not set quickcheck informs about the
 | 
1610  | 
current size for assignment values.  | 
|
| 
40245
 
59f011c1877a
updating documentation on quickcheck in the Isar reference
 
bulwahn 
parents: 
40171 
diff
changeset
 | 
1611  | 
|
| 40254 | 1612  | 
    \item[@{text expect}] can be used to check if the user's
 | 
1613  | 
    expectation was met (@{text no_expectation}, @{text
 | 
|
1614  | 
    no_counterexample}, or @{text counterexample}).
 | 
|
| 
40245
 
59f011c1877a
updating documentation on quickcheck in the Isar reference
 
bulwahn 
parents: 
40171 
diff
changeset
 | 
1615  | 
|
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1616  | 
    \end{description}
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1617  | 
|
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1618  | 
These option can be given within square brackets.  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1619  | 
|
| 
42215
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1620  | 
  \item @{command (HOL) "quickcheck_params"} changes
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1621  | 
    @{command (HOL) "quickcheck"} configuration options persistently.
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1622  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1623  | 
  \item @{command (HOL) "refute"} tests the current goal for
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1624  | 
counterexamples using a reduction to SAT. The following configuration  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1625  | 
options are supported:  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1626  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1627  | 
    \begin{description}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1628  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1629  | 
    \item[@{text minsize}] specifies the minimum size (cardinality) of the
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1630  | 
models to search for.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1631  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1632  | 
    \item[@{text maxsize}] specifies the maximum size (cardinality) of the
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1633  | 
models to search for. Nonpositive values mean $\infty$.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1634  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1635  | 
    \item[@{text maxvars}] specifies the maximum number of Boolean variables
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1636  | 
to use when transforming the term into a propositional formula.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1637  | 
Nonpositive values mean $\infty$.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1638  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1639  | 
    \item[@{text satsolver}] specifies the SAT solver to use.
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1640  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1641  | 
    \item[@{text no_assms}] specifies whether assumptions in
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1642  | 
structured proofs should be ignored.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1643  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1644  | 
    \item[@{text maxtime}] sets the time limit in seconds.
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1645  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1646  | 
    \item[@{text expect}] can be used to check if the user's
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1647  | 
    expectation was met (@{text genuine}, @{text potential},
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1648  | 
    @{text none}, or @{text unknown}).
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1649  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1650  | 
    \end{description}
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1651  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1652  | 
These option can be given within square brackets.  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1653  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1654  | 
  \item @{command (HOL) "refute_params"} changes
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1655  | 
    @{command (HOL) "refute"} configuration options persistently.
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1656  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1657  | 
  \item @{command (HOL) "nitpick"} tests the current goal for counterexamples
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1658  | 
using a reduction to first-order relational logic. See the Nitpick manual  | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1659  | 
    \cite{isabelle-nitpick} for details.
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1660  | 
|
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1661  | 
  \item @{command (HOL) "nitpick_params"} changes
 | 
| 
 
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
 
blanchet 
parents: 
42123 
diff
changeset
 | 
1662  | 
    @{command (HOL) "nitpick"} configuration options persistently.
 | 
| 
31912
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1663  | 
|
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1664  | 
  \end{description}
 | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1665  | 
*}  | 
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1666  | 
|
| 
 
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
 
haftmann 
parents: 
31254 
diff
changeset
 | 
1667  | 
|
| 28752 | 1668  | 
section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *}
 | 
| 26849 | 1669  | 
|
1670  | 
text {*
 | 
|
| 
27123
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1671  | 
The following tools of Isabelle/HOL support cases analysis and  | 
| 
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1672  | 
induction in unstructured tactic scripts; see also  | 
| 
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1673  | 
  \secref{sec:cases-induct} for proper Isar versions of similar ideas.
 | 
| 26849 | 1674  | 
|
1675  | 
  \begin{matharray}{rcl}
 | 
|
| 
28761
 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 
wenzelm 
parents: 
28760 
diff
changeset
 | 
1676  | 
    @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 
wenzelm 
parents: 
28760 
diff
changeset
 | 
1677  | 
    @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 
wenzelm 
parents: 
28760 
diff
changeset
 | 
1678  | 
    @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 
wenzelm 
parents: 
28760 
diff
changeset
 | 
1679  | 
    @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
| 26849 | 1680  | 
  \end{matharray}
 | 
1681  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1682  | 
  @{rail "
 | 
| 42705 | 1683  | 
    @@{method (HOL) case_tac} @{syntax goal_spec}? @{syntax term} rule?
 | 
| 26849 | 1684  | 
;  | 
| 42705 | 1685  | 
    @@{method (HOL) induct_tac} @{syntax goal_spec}? (@{syntax insts} * @'and') rule?
 | 
| 26849 | 1686  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1687  | 
    @@{method (HOL) ind_cases} (@{syntax prop}+) (@'for' (@{syntax name}+))?
 | 
| 26849 | 1688  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1689  | 
    @@{command (HOL) inductive_cases} (@{syntax thmdecl}? (@{syntax prop}+) + @'and')
 | 
| 26849 | 1690  | 
;  | 
1691  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1692  | 
    rule: 'rule' ':' @{syntax thmref}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1693  | 
"}  | 
| 26849 | 1694  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1695  | 
  \begin{description}
 | 
| 26849 | 1696  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1697  | 
  \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit
 | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1698  | 
to reason about inductive types. Rules are selected according to  | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1699  | 
  the declarations by the @{attribute cases} and @{attribute induct}
 | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1700  | 
  attributes, cf.\ \secref{sec:cases-induct}.  The @{command (HOL)
 | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1701  | 
datatype} package already takes care of this.  | 
| 
27123
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1702  | 
|
| 
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1703  | 
These unstructured tactics feature both goal addressing and dynamic  | 
| 26849 | 1704  | 
  instantiation.  Note that named rule cases are \emph{not} provided
 | 
| 
27123
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1705  | 
  as would be by the proper @{method cases} and @{method induct} proof
 | 
| 
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1706  | 
  methods (see \secref{sec:cases-induct}).  Unlike the @{method
 | 
| 
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1707  | 
  induct} method, @{method induct_tac} does not handle structured rule
 | 
| 
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1708  | 
statements, only the compact object-logic conclusion of the subgoal  | 
| 
 
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
 
wenzelm 
parents: 
27103 
diff
changeset
 | 
1709  | 
being addressed.  | 
| 42123 | 1710  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1711  | 
  \item @{method (HOL) ind_cases} and @{command (HOL)
 | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1712  | 
  "inductive_cases"} provide an interface to the internal @{ML_text
 | 
| 26860 | 1713  | 
mk_cases} operation. Rules are simplified in an unrestricted  | 
1714  | 
forward manner.  | 
|
| 26849 | 1715  | 
|
1716  | 
  While @{method (HOL) ind_cases} is a proof method to apply the
 | 
|
1717  | 
  result immediately as elimination rules, @{command (HOL)
 | 
|
1718  | 
"inductive_cases"} provides case split theorems at the theory level  | 
|
1719  | 
  for later use.  The @{keyword "for"} argument of the @{method (HOL)
 | 
|
1720  | 
ind_cases} method allows to specify a list of variables that should  | 
|
1721  | 
be generalized before applying the resulting rule.  | 
|
1722  | 
||
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
1723  | 
  \end{description}
 | 
| 26849 | 1724  | 
*}  | 
1725  | 
||
1726  | 
||
1727  | 
section {* Executable code *}
 | 
|
1728  | 
||
| 
42627
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1729  | 
text {* For validation purposes, it is often useful to \emph{execute}
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1730  | 
specifications. In principle, execution could be simulated by  | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1731  | 
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
 | 
1732  | 
simplification. Unfortunately, this approach is rather inefficient.  | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1733  | 
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
 | 
1734  | 
them into a functional programming language such as ML.  | 
| 26849 | 1735  | 
|
| 45192 | 1736  | 
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
 | 
1737  | 
from executable specifications. Isabelle/HOL instantiates these  | 
| 45192 | 1738  | 
mechanisms in a way that is amenable to end-user applications. Code  | 
1739  | 
can be generated for functional programs (including overloading  | 
|
1740  | 
  using type classes) targeting SML \cite{SML}, OCaml \cite{OCaml},
 | 
|
1741  | 
  Haskell \cite{haskell-revised-report} and Scala
 | 
|
| 
42627
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1742  | 
  \cite{scala-overview-tech-report}.  Conceptually, code generation is
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1743  | 
  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
 | 
1744  | 
  \emph{translation} into an abstract executable view and
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1745  | 
  \emph{serialization} to a specific \emph{target language}.
 | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1746  | 
Inductive specifications can be executed using the predicate  | 
| 
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1747  | 
  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
 | 
1748  | 
an introduction.  | 
| 37422 | 1749  | 
|
1750  | 
  \begin{matharray}{rcl}
 | 
|
1751  | 
    @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
|
1752  | 
    @{attribute_def (HOL) code} & : & @{text attribute} \\
 | 
|
1753  | 
    @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
1754  | 
    @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
1755  | 
    @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
|
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1756  | 
    @{attribute_def (HOL) code_unfold} & : & @{text attribute} \\
 | 
| 37422 | 1757  | 
    @{attribute_def (HOL) code_post} & : & @{text attribute} \\
 | 
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1758  | 
    @{attribute_def (HOL) code_unfold_post} & : & @{text attribute} \\
 | 
| 37422 | 1759  | 
    @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
1760  | 
    @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
|
1761  | 
    @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
|
1762  | 
    @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
1763  | 
    @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
1764  | 
    @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
1765  | 
    @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
1766  | 
    @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
1767  | 
    @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
1768  | 
    @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
1769  | 
    @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
| 
45408
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1770  | 
    @{command_def (HOL) "code_reflect"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1771  | 
    @{command_def (HOL) "code_pred"} & : & @{text "theory \<rightarrow> proof(prove)"}
 | 
| 37422 | 1772  | 
  \end{matharray}
 | 
1773  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1774  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1775  | 
    @@{command (HOL) export_code} ( constexpr + ) \\
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1776  | 
       ( ( @'in' target ( @'module_name' @{syntax string} ) ? \\
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1777  | 
        ( @'file' ( @{syntax string} | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
 | 
| 37422 | 1778  | 
;  | 
1779  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1780  | 
    const: @{syntax term}
 | 
| 37422 | 1781  | 
;  | 
1782  | 
||
| 
40711
 
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
 
haftmann 
parents: 
40709 
diff
changeset
 | 
1783  | 
constexpr: ( const | 'name._' | '_' )  | 
| 37422 | 1784  | 
;  | 
1785  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1786  | 
    typeconstructor: @{syntax nameref}
 | 
| 37422 | 1787  | 
;  | 
1788  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1789  | 
    class: @{syntax nameref}
 | 
| 37422 | 1790  | 
;  | 
1791  | 
||
| 38814 | 1792  | 
target: 'SML' | 'OCaml' | 'Haskell' | 'Scala'  | 
| 37422 | 1793  | 
;  | 
1794  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1795  | 
    @@{attribute (HOL) code} ( 'del' | 'abstype' | 'abstract' )?
 | 
| 37422 | 1796  | 
;  | 
1797  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1798  | 
    @@{command (HOL) code_abort} ( const + )
 | 
| 37422 | 1799  | 
;  | 
1800  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1801  | 
    @@{command (HOL) code_datatype} ( const + )
 | 
| 37422 | 1802  | 
;  | 
1803  | 
||
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1804  | 
    @@{attribute (HOL) code_unfold} ( 'del' ) ?
 | 
| 37422 | 1805  | 
;  | 
1806  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1807  | 
    @@{attribute (HOL) code_post} ( 'del' ) ?
 | 
| 37422 | 1808  | 
;  | 
1809  | 
||
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1810  | 
    @@{attribute (HOL) code_unfold_post}
 | 
| 
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1811  | 
;  | 
| 
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1812  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1813  | 
    @@{command (HOL) code_thms} ( constexpr + ) ?
 | 
| 37422 | 1814  | 
;  | 
1815  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1816  | 
    @@{command (HOL) code_deps} ( constexpr + ) ?
 | 
| 37422 | 1817  | 
;  | 
1818  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1819  | 
    @@{command (HOL) code_const} (const + @'and') \\
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1820  | 
      ( ( '(' target ( syntax ? + @'and' ) ')' ) + )
 | 
| 37422 | 1821  | 
;  | 
1822  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1823  | 
    @@{command (HOL) code_type} (typeconstructor + @'and') \\
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1824  | 
      ( ( '(' target ( syntax ? + @'and' ) ')' ) + )
 | 
| 37422 | 1825  | 
;  | 
1826  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1827  | 
    @@{command (HOL) code_class} (class + @'and') \\
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1828  | 
      ( ( '(' target \\ ( @{syntax string} ? + @'and' ) ')' ) + )
 | 
| 37422 | 1829  | 
;  | 
1830  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1831  | 
    @@{command (HOL) code_instance} (( typeconstructor '::' class ) + @'and') \\
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1832  | 
      ( ( '(' target ( '-' ? + @'and' ) ')' ) + )
 | 
| 37422 | 1833  | 
;  | 
1834  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1835  | 
    @@{command (HOL) code_reserved} target ( @{syntax string} + )
 | 
| 37422 | 1836  | 
;  | 
1837  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1838  | 
    @@{command (HOL) code_monad} const const target
 | 
| 37422 | 1839  | 
;  | 
1840  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1841  | 
    @@{command (HOL) code_include} target ( @{syntax string} ( @{syntax string} | '-') )
 | 
| 37422 | 1842  | 
;  | 
1843  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1844  | 
    @@{command (HOL) code_modulename} target ( ( @{syntax string} @{syntax string} ) + )
 | 
| 39608 | 1845  | 
;  | 
1846  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1847  | 
    @@{command (HOL) code_reflect} @{syntax string} \\
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1848  | 
      ( @'datatypes' ( @{syntax string} '=' ( '_' | ( @{syntax string} + '|' ) + @'and' ) ) ) ? \\
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1849  | 
      ( @'functions' ( @{syntax string} + ) ) ? ( @'file' @{syntax string} ) ?
 | 
| 37422 | 1850  | 
;  | 
1851  | 
||
| 
45408
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1852  | 
    @@{command (HOL) code_pred} \\( '(' @'modes' ':' modedecl ')')? \\ const
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1853  | 
;  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1854  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1855  | 
    syntax: @{syntax string} | ( @'infix' | @'infixl' | @'infixr' ) @{syntax nat} @{syntax string}
 | 
| 
45408
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1856  | 
;  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1857  | 
|
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1858  | 
modedecl: (modes | ((const ':' modes) \\  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1859  | 
(@'and' ((const ':' modes @'and') +))?))  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1860  | 
;  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1861  | 
|
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1862  | 
modes: mode @'as' const  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1863  | 
"}  | 
| 37422 | 1864  | 
|
1865  | 
  \begin{description}
 | 
|
1866  | 
||
1867  | 
  \item @{command (HOL) "export_code"} generates code for a given list
 | 
|
| 39608 | 1868  | 
of constants in the specified target language(s). If no  | 
1869  | 
serialization instruction is given, only abstract code is generated  | 
|
1870  | 
internally.  | 
|
| 37422 | 1871  | 
|
1872  | 
Constants may be specified by giving them literally, referring to  | 
|
1873  | 
  all executable contants within a certain theory by giving @{text
 | 
|
1874  | 
  "name.*"}, or referring to \emph{all} executable constants currently
 | 
|
1875  | 
  available by giving @{text "*"}.
 | 
|
1876  | 
||
1877  | 
By default, for each involved theory one corresponding name space  | 
|
1878  | 
module is generated. Alternativly, a module name may be specified  | 
|
1879  | 
  after the @{keyword "module_name"} keyword; then \emph{all} code is
 | 
|
1880  | 
placed in this module.  | 
|
1881  | 
||
| 39608 | 1882  | 
  For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification
 | 
1883  | 
  refers to a single file; for \emph{Haskell}, it refers to a whole
 | 
|
1884  | 
directory, where code is generated in multiple files reflecting the  | 
|
1885  | 
module hierarchy. Omitting the file specification denotes standard  | 
|
| 37749 | 1886  | 
output.  | 
| 37422 | 1887  | 
|
1888  | 
Serializers take an optional list of arguments in parentheses. For  | 
|
1889  | 
  \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits
 | 
|
1890  | 
explicit module signatures.  | 
|
| 42123 | 1891  | 
|
| 39608 | 1892  | 
  For \emph{Haskell} a module name prefix may be given using the
 | 
1893  | 
  ``@{text "root:"}'' argument; ``@{text string_classes}'' adds a
 | 
|
1894  | 
  ``@{verbatim "deriving (Read, Show)"}'' clause to each appropriate
 | 
|
1895  | 
datatype declaration.  | 
|
| 37422 | 1896  | 
|
1897  | 
  \item @{attribute (HOL) code} explicitly selects (or with option
 | 
|
| 
38462
 
34d3de1254cd
formally document `code abstype` and `code abstract` attributes
 
haftmann 
parents: 
37820 
diff
changeset
 | 
1898  | 
  ``@{text "del"}'' deselects) a code equation for code generation.
 | 
| 
 
34d3de1254cd
formally document `code abstype` and `code abstract` attributes
 
haftmann 
parents: 
37820 
diff
changeset
 | 
1899  | 
Usually packages introducing code equations provide a reasonable  | 
| 
 
34d3de1254cd
formally document `code abstype` and `code abstract` attributes
 
haftmann 
parents: 
37820 
diff
changeset
 | 
1900  | 
  default setup for selection.  Variants @{text "code abstype"} and
 | 
| 
 
34d3de1254cd
formally document `code abstype` and `code abstract` attributes
 
haftmann 
parents: 
37820 
diff
changeset
 | 
1901  | 
  @{text "code abstract"} declare abstract datatype certificates or
 | 
| 
 
34d3de1254cd
formally document `code abstype` and `code abstract` attributes
 
haftmann 
parents: 
37820 
diff
changeset
 | 
1902  | 
code equations on abstract datatype representations respectively.  | 
| 37422 | 1903  | 
|
1904  | 
  \item @{command (HOL) "code_abort"} declares constants which are not
 | 
|
| 39608 | 1905  | 
required to have a definition by means of code equations; if needed  | 
1906  | 
these are implemented by program abort instead.  | 
|
| 37422 | 1907  | 
|
1908  | 
  \item @{command (HOL) "code_datatype"} specifies a constructor set
 | 
|
1909  | 
for a logical type.  | 
|
1910  | 
||
1911  | 
  \item @{command (HOL) "print_codesetup"} gives an overview on
 | 
|
1912  | 
selected code equations and code generator datatypes.  | 
|
1913  | 
||
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1914  | 
  \item @{attribute (HOL) code_unfold} declares (or with option
 | 
| 
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1915  | 
  ``@{text "del"}'' removes) theorems which are applied as
 | 
| 39608 | 1916  | 
rewrite rules to any code equation during preprocessing.  | 
| 37422 | 1917  | 
|
| 39608 | 1918  | 
  \item @{attribute (HOL) code_post} declares (or with option ``@{text
 | 
1919  | 
"del"}'' removes) theorems which are applied as rewrite rules to any  | 
|
1920  | 
result of an evaluation.  | 
|
| 37422 | 1921  | 
|
| 
45232
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1922  | 
  \item @{attribute (HOL) code_unfold_post} declares equations which are
 | 
| 
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1923  | 
applied as rewrite rules to any code equation during preprocessing,  | 
| 
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1924  | 
and symmetrically to any result of an evaluation.  | 
| 
 
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
 
bulwahn 
parents: 
45192 
diff
changeset
 | 
1925  | 
|
| 39608 | 1926  | 
  \item @{command (HOL) "print_codeproc"} prints the setup of the code
 | 
1927  | 
generator preprocessor.  | 
|
| 37422 | 1928  | 
|
1929  | 
  \item @{command (HOL) "code_thms"} prints a list of theorems
 | 
|
1930  | 
representing the corresponding program containing all given  | 
|
1931  | 
constants after preprocessing.  | 
|
1932  | 
||
1933  | 
  \item @{command (HOL) "code_deps"} visualizes dependencies of
 | 
|
1934  | 
theorems representing the corresponding program containing all given  | 
|
1935  | 
constants after preprocessing.  | 
|
1936  | 
||
1937  | 
  \item @{command (HOL) "code_const"} associates a list of constants
 | 
|
1938  | 
with target-specific serializations; omitting a serialization  | 
|
1939  | 
deletes an existing serialization.  | 
|
1940  | 
||
1941  | 
  \item @{command (HOL) "code_type"} associates a list of type
 | 
|
1942  | 
constructors with target-specific serializations; omitting a  | 
|
1943  | 
serialization deletes an existing serialization.  | 
|
1944  | 
||
1945  | 
  \item @{command (HOL) "code_class"} associates a list of classes
 | 
|
1946  | 
with target-specific class names; omitting a serialization deletes  | 
|
1947  | 
  an existing serialization.  This applies only to \emph{Haskell}.
 | 
|
1948  | 
||
1949  | 
  \item @{command (HOL) "code_instance"} declares a list of type
 | 
|
1950  | 
constructor / class instance relations as ``already present'' for a  | 
|
1951  | 
  given target.  Omitting a ``@{text "-"}'' deletes an existing
 | 
|
1952  | 
``already present'' declaration. This applies only to  | 
|
1953  | 
  \emph{Haskell}.
 | 
|
1954  | 
||
1955  | 
  \item @{command (HOL) "code_reserved"} declares a list of names as
 | 
|
1956  | 
reserved for a given target, preventing it to be shadowed by any  | 
|
1957  | 
generated code.  | 
|
1958  | 
||
1959  | 
  \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
 | 
|
1960  | 
to generate monadic code for Haskell.  | 
|
1961  | 
||
1962  | 
  \item @{command (HOL) "code_include"} adds arbitrary named content
 | 
|
1963  | 
  (``include'') to generated code.  A ``@{text "-"}'' as last argument
 | 
|
1964  | 
will remove an already added ``include''.  | 
|
1965  | 
||
1966  | 
  \item @{command (HOL) "code_modulename"} declares aliasings from one
 | 
|
1967  | 
module name onto another.  | 
|
1968  | 
||
| 39608 | 1969  | 
  \item @{command (HOL) "code_reflect"} without a ``@{text "file"}''
 | 
1970  | 
argument compiles code into the system runtime environment and  | 
|
1971  | 
modifies the code generator setup that future invocations of system  | 
|
1972  | 
  runtime code generation referring to one of the ``@{text
 | 
|
1973  | 
  "datatypes"}'' or ``@{text "functions"}'' entities use these precompiled
 | 
|
1974  | 
  entities.  With a ``@{text "file"}'' argument, the corresponding code
 | 
|
1975  | 
is generated into that specified file without modifying the code  | 
|
1976  | 
generator setup.  | 
|
1977  | 
||
| 
45408
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1978  | 
  \item @{command (HOL) "code_pred"} creates code equations for a predicate
 | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1979  | 
given a set of introduction rules. Optional mode annotations determine  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1980  | 
which arguments are supposed to be input or output. If alternative  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1981  | 
introduction rules are declared, one must prove a corresponding elimination  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1982  | 
rule.  | 
| 
 
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
 
bulwahn 
parents: 
45232 
diff
changeset
 | 
1983  | 
|
| 37422 | 1984  | 
  \end{description}
 | 
| 
42627
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1985  | 
*}  | 
| 37422 | 1986  | 
|
| 
42627
 
8749742785b8
moved material about old codegen to isar-ref manual;
 
wenzelm 
parents: 
42626 
diff
changeset
 | 
1987  | 
|
| 27045 | 1988  | 
section {* Definition by specification \label{sec:hol-specification} *}
 | 
1989  | 
||
1990  | 
text {*
 | 
|
1991  | 
  \begin{matharray}{rcl}
 | 
|
| 
28761
 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 
wenzelm 
parents: 
28760 
diff
changeset
 | 
1992  | 
    @{command_def (HOL) "specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
 | 
| 
 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 
wenzelm 
parents: 
28760 
diff
changeset
 | 
1993  | 
    @{command_def (HOL) "ax_specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
 | 
| 27045 | 1994  | 
  \end{matharray}
 | 
1995  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1996  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1997  | 
  (@@{command (HOL) specification} | @@{command (HOL) ax_specification})
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
1998  | 
    '(' (decl +) ')' \\ (@{syntax thmdecl}? @{syntax prop} +)
 | 
| 27045 | 1999  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2000  | 
  decl: ((@{syntax name} ':')? @{syntax term} '(' @'overloaded' ')'?)
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42215 
diff
changeset
 | 
2001  | 
"}  | 
| 27045 | 2002  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
2003  | 
  \begin{description}
 | 
| 27045 | 2004  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
2005  | 
  \item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a
 | 
| 27045 | 2006  | 
goal stating the existence of terms with the properties specified to  | 
2007  | 
  hold for the constants given in @{text decls}.  After finishing the
 | 
|
2008  | 
proof, the theory will be augmented with definitions for the given  | 
|
2009  | 
constants, as well as with theorems stating the properties for these  | 
|
2010  | 
constants.  | 
|
2011  | 
||
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
2012  | 
  \item @{command (HOL) "ax_specification"}~@{text "decls \<phi>"} sets up
 | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
2013  | 
a goal stating the existence of terms with the properties specified  | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
2014  | 
  to hold for the constants given in @{text decls}.  After finishing
 | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
2015  | 
the proof, the theory will be augmented with axioms expressing the  | 
| 
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
2016  | 
properties given in the first place.  | 
| 27045 | 2017  | 
|
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
2018  | 
  \item @{text decl} declares a constant to be defined by the
 | 
| 27045 | 2019  | 
  specification given.  The definition for the constant @{text c} is
 | 
2020  | 
  bound to the name @{text c_def} unless a theorem name is given in
 | 
|
2021  | 
the declaration. Overloaded constants should be declared as such.  | 
|
2022  | 
||
| 
28760
 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 
wenzelm 
parents: 
28752 
diff
changeset
 | 
2023  | 
  \end{description}
 | 
| 27045 | 2024  | 
|
2025  | 
  Whether to use @{command (HOL) "specification"} or @{command (HOL)
 | 
|
2026  | 
  "ax_specification"} is to some extent a matter of style.  @{command
 | 
|
2027  | 
(HOL) "specification"} introduces no new axioms, and so by  | 
|
2028  | 
  construction cannot introduce inconsistencies, whereas @{command
 | 
|
2029  | 
(HOL) "ax_specification"} does introduce axioms, but only after the  | 
|
2030  | 
user has explicitly proven it to be safe. A practical issue must be  | 
|
2031  | 
considered, though: After introducing two constants with the same  | 
|
2032  | 
  properties using @{command (HOL) "specification"}, one can prove
 | 
|
2033  | 
that the two constants are, in fact, equal. If this might be a  | 
|
2034  | 
  problem, one should use @{command (HOL) "ax_specification"}.
 | 
|
2035  | 
*}  | 
|
2036  | 
||
| 26840 | 2037  | 
end  |