author | blanchet |
Thu, 11 Sep 2014 18:54:36 +0200 | |
changeset 58306 | 117ba6cbe414 |
parent 58305 | 57752a91eec4 |
child 58310 | 91ea607a34d8 |
permissions | -rw-r--r-- |
26840 | 1 |
theory HOL_Specific |
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
2 |
imports Base Main "~~/src/HOL/Library/Old_Recdef" "~~/src/Tools/Adhoc_Overloading" |
26840 | 3 |
begin |
4 |
||
50109 | 5 |
chapter {* Higher-Order Logic *} |
42914 | 6 |
|
7 |
text {* Isabelle/HOL is based on Higher-Order Logic, a polymorphic |
|
8 |
version of Church's Simple Theory of Types. HOL can be best |
|
9 |
understood as a simply-typed version of classical set theory. The |
|
10 |
logic was first implemented in Gordon's HOL system |
|
11 |
\cite{mgordon-hol}. It extends Church's original logic |
|
12 |
\cite{church40} by explicit type variables (naive polymorphism) and |
|
13 |
a sound axiomatization scheme for new types based on subsets of |
|
14 |
existing types. |
|
15 |
||
16 |
Andrews's book \cite{andrews86} is a full description of the |
|
17 |
original Church-style higher-order logic, with proofs of correctness |
|
18 |
and completeness wrt.\ certain set-theoretic interpretations. The |
|
19 |
particular extensions of Gordon-style HOL are explained semantically |
|
20 |
in two chapters of the 1993 HOL book \cite{pitts93}. |
|
21 |
||
22 |
Experience with HOL over decades has demonstrated that higher-order |
|
23 |
logic is widely applicable in many areas of mathematics and computer |
|
24 |
science. In a sense, Higher-Order Logic is simpler than First-Order |
|
25 |
Logic, because there are fewer restrictions and special cases. Note |
|
26 |
that HOL is \emph{weaker} than FOL with axioms for ZF set theory, |
|
27 |
which is traditionally considered the standard foundation of regular |
|
28 |
mathematics, but for most applications this does not matter. If you |
|
29 |
prefer ML to Lisp, you will probably prefer HOL to ZF. |
|
30 |
||
31 |
\medskip The syntax of HOL follows @{text "\<lambda>"}-calculus and |
|
32 |
functional programming. Function application is curried. To apply |
|
33 |
the function @{text f} of type @{text "\<tau>\<^sub>1 \<Rightarrow> \<tau>\<^sub>2 \<Rightarrow> \<tau>\<^sub>3"} to the |
|
34 |
arguments @{text a} and @{text b} in HOL, you simply write @{text "f |
|
35 |
a b"} (as in ML or Haskell). There is no ``apply'' operator; the |
|
36 |
existing application of the Pure @{text "\<lambda>"}-calculus is re-used. |
|
37 |
Note that in HOL @{text "f (a, b)"} means ``@{text "f"} applied to |
|
38 |
the pair @{text "(a, b)"} (which is notation for @{text "Pair a |
|
39 |
b"}). The latter typically introduces extra formal efforts that can |
|
40 |
be avoided by currying functions by default. Explicit tuples are as |
|
41 |
infrequent in HOL formalizations as in good ML or Haskell programs. |
|
42 |
||
43 |
\medskip Isabelle/HOL has a distinct feel, compared to other |
|
44 |
object-logics like Isabelle/ZF. It identifies object-level types |
|
45 |
with meta-level types, taking advantage of the default |
|
46 |
type-inference mechanism of Isabelle/Pure. HOL fully identifies |
|
47 |
object-level functions with meta-level functions, with native |
|
48 |
abstraction and application. |
|
49 |
||
50 |
These identifications allow Isabelle to support HOL particularly |
|
51 |
nicely, but they also mean that HOL requires some sophistication |
|
52 |
from the user. In particular, an understanding of Hindley-Milner |
|
53 |
type-inference with type-classes, which are both used extensively in |
|
54 |
the standard libraries and applications. Beginners can set |
|
55 |
@{attribute show_types} or even @{attribute show_sorts} to get more |
|
56 |
explicit information about the result of type-inference. *} |
|
57 |
||
58 |
||
50109 | 59 |
chapter {* Derived specification elements *} |
60 |
||
42908 | 61 |
section {* Inductive and coinductive definitions \label{sec:hol-inductive} *} |
62 |
||
46280 | 63 |
text {* |
64 |
\begin{matharray}{rcl} |
|
65 |
@{command_def (HOL) "inductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
|
66 |
@{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
|
67 |
@{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
|
68 |
@{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
|
50302 | 69 |
@{command_def "print_inductives"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
46280 | 70 |
@{attribute_def (HOL) mono} & : & @{text attribute} \\ |
71 |
\end{matharray} |
|
72 |
||
73 |
An \emph{inductive definition} specifies the least predicate or set |
|
74 |
@{text R} closed under given rules: applying a rule to elements of |
|
75 |
@{text R} yields a result within @{text R}. For example, a |
|
76 |
structural operational semantics is an inductive definition of an |
|
77 |
evaluation relation. |
|
42908 | 78 |
|
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
79 |
Dually, a \emph{coinductive definition} specifies the greatest |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
80 |
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
|
81 |
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
|
82 |
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
|
83 |
bisimulation relations to formalise equivalence of processes and |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
84 |
infinite data structures. |
47859 | 85 |
|
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
86 |
Both inductive and coinductive definitions are based on the |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
87 |
Knaster-Tarski fixed-point theorem for complete lattices. The |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
88 |
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
|
89 |
functor on subsets of set-theoretic relations. The required |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
90 |
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
|
91 |
the fixed-point definition and the resulting consequences. This |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
92 |
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
|
93 |
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
|
94 |
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
|
95 |
@{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
|
96 |
introduction rule. The default rule declarations of Isabelle/HOL |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
97 |
already take care of most common situations. |
42907
dfd4ef8e73f6
updated and re-unified HOL typedef, with some live examples;
wenzelm
parents:
42705
diff
changeset
|
98 |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
99 |
@{rail \<open> |
42908 | 100 |
(@@{command (HOL) inductive} | @@{command (HOL) inductive_set} | |
101 |
@@{command (HOL) coinductive} | @@{command (HOL) coinductive_set}) |
|
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
102 |
@{syntax target}? \<newline> |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
103 |
@{syntax "fixes"} (@'for' @{syntax "fixes"})? (@'where' clauses)? \<newline> |
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
104 |
(@'monos' @{syntax thmrefs})? |
42908 | 105 |
; |
106 |
clauses: (@{syntax thmdecl}? @{syntax prop} + '|') |
|
107 |
; |
|
108 |
@@{attribute (HOL) mono} (() | 'add' | 'del') |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
109 |
\<close>} |
42908 | 110 |
|
111 |
\begin{description} |
|
112 |
||
113 |
\item @{command (HOL) "inductive"} and @{command (HOL) |
|
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
114 |
"coinductive"} define (co)inductive predicates from the introduction |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
115 |
rules. |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
116 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
117 |
The propositions given as @{text "clauses"} in the @{keyword |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
118 |
"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
|
119 |
(with arbitrary nesting), or equalities using @{text "\<equiv>"}. The |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
120 |
latter specifies extra-logical abbreviations in the sense of |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
121 |
@{command_ref abbreviation}. Introducing abstract syntax |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
122 |
simultaneously with the actual introduction rules is occasionally |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
123 |
useful for complex specifications. |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
124 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
125 |
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
|
126 |
the (co)inductive predicates that remain fixed throughout the |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
127 |
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
|
128 |
in each occurrence within the given @{text "clauses"}. |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
129 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
130 |
The optional @{keyword "monos"} declaration contains additional |
42908 | 131 |
\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
|
132 |
applied to a recursive set in the introduction rules. |
42908 | 133 |
|
134 |
\item @{command (HOL) "inductive_set"} and @{command (HOL) |
|
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
135 |
"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
|
136 |
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
|
137 |
where multiple arguments are simulated via tuples. |
42908 | 138 |
|
50302 | 139 |
\item @{command "print_inductives"} prints (co)inductive definitions and |
140 |
monotonicity rules. |
|
141 |
||
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
142 |
\item @{attribute (HOL) mono} declares monotonicity rules in the |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
143 |
context. These rule are involved in the automated monotonicity |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
144 |
proof of the above inductive and coinductive definitions. |
42908 | 145 |
|
146 |
\end{description} |
|
147 |
*} |
|
148 |
||
149 |
||
150 |
subsection {* Derived rules *} |
|
151 |
||
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
152 |
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
|
153 |
main theorems: |
42908 | 154 |
|
155 |
\begin{description} |
|
156 |
||
157 |
\item @{text R.intros} is the list of introduction rules as proven |
|
158 |
theorems, for the recursive predicates (or sets). The rules are |
|
159 |
also available individually, using the names given them in the |
|
160 |
theory file; |
|
161 |
||
162 |
\item @{text R.cases} is the case analysis (or elimination) rule; |
|
163 |
||
164 |
\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
|
165 |
rule; |
336752fb25df
adding documentation about simps equation in the inductive package
bulwahn
parents:
44055
diff
changeset
|
166 |
|
336752fb25df
adding documentation about simps equation in the inductive package
bulwahn
parents:
44055
diff
changeset
|
167 |
\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
|
168 |
predicate one step. |
47859 | 169 |
|
42908 | 170 |
\end{description} |
171 |
||
172 |
When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are |
|
173 |
defined simultaneously, the list of introduction rules is called |
|
174 |
@{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are |
|
175 |
called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list |
|
176 |
of mutual induction rules is called @{text |
|
177 |
"R\<^sub>1_\<dots>_R\<^sub>n.inducts"}. |
|
178 |
*} |
|
179 |
||
180 |
||
181 |
subsection {* Monotonicity theorems *} |
|
182 |
||
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
183 |
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
|
184 |
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
|
185 |
@{attribute (HOL) mono} attribute. See the main Isabelle/HOL |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
186 |
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
|
187 |
theorems is as follows: |
42908 | 188 |
|
189 |
\begin{itemize} |
|
190 |
||
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
191 |
\item Theorems of the form @{text "A \<le> B \<Longrightarrow> \<M> A \<le> \<M> B"}, for proving |
42908 | 192 |
monotonicity of inductive definitions whose introduction rules have |
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
193 |
premises involving terms such as @{text "\<M> R t"}. |
42908 | 194 |
|
195 |
\item Monotonicity theorems for logical operators, which are of the |
|
196 |
general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}. For example, in |
|
197 |
the case of the operator @{text "\<or>"}, the corresponding theorem is |
|
198 |
\[ |
|
199 |
\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"}} |
|
200 |
\] |
|
201 |
||
202 |
\item De Morgan style equations for reasoning about the ``polarity'' |
|
203 |
of expressions, e.g. |
|
204 |
\[ |
|
205 |
@{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad |
|
206 |
@{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"} |
|
207 |
\] |
|
208 |
||
209 |
\item Equations for reducing complex operators to more primitive |
|
210 |
ones whose monotonicity can easily be proved, e.g. |
|
211 |
\[ |
|
212 |
@{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad |
|
213 |
@{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"} |
|
214 |
\] |
|
215 |
||
216 |
\end{itemize} |
|
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
217 |
*} |
42908 | 218 |
|
42913
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
219 |
subsubsection {* Examples *} |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
220 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
221 |
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
|
222 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
223 |
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
|
224 |
where |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
225 |
empty: "{} \<in> Fin A" |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
226 |
| 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
|
227 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
228 |
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
|
229 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
230 |
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
|
231 |
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
|
232 |
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
|
233 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
234 |
text {* Common logical connectives can be easily characterized as |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
235 |
non-recursive inductive definitions with parameters, but without |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
236 |
arguments. *} |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
237 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
238 |
inductive AND for A B :: bool |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
239 |
where "A \<Longrightarrow> B \<Longrightarrow> AND A B" |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
240 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
241 |
inductive OR for A B :: bool |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
242 |
where "A \<Longrightarrow> OR A B" |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
243 |
| "B \<Longrightarrow> OR A B" |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
244 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
245 |
inductive EXISTS for B :: "'a \<Rightarrow> bool" |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
246 |
where "B a \<Longrightarrow> EXISTS B" |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
247 |
|
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
248 |
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
|
249 |
the @{command inductive} package coincide with the expected |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
250 |
elimination rules for Natural Deduction. Already in the original |
68bc69bdce88
updated and re-unified (co)inductive definitions in HOL;
wenzelm
parents:
42912
diff
changeset
|
251 |
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
|
252 |
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
|
253 |
elimination can be constructed systematically. *} |
42908 | 254 |
|
255 |
||
256 |
section {* Recursive functions \label{sec:recursion} *} |
|
257 |
||
258 |
text {* |
|
259 |
\begin{matharray}{rcl} |
|
260 |
@{command_def (HOL) "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
|
261 |
@{command_def (HOL) "fun"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
|
262 |
@{command_def (HOL) "function"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\ |
|
263 |
@{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\ |
|
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53982
diff
changeset
|
264 |
@{command_def (HOL) "fun_cases"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
42908 | 265 |
\end{matharray} |
266 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
267 |
@{rail \<open> |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
268 |
@@{command (HOL) primrec} @{syntax target}? @{syntax "fixes"} @'where' equations |
42908 | 269 |
; |
270 |
(@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts? |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
271 |
@{syntax "fixes"} \<newline> @'where' equations |
26849 | 272 |
; |
273 |
||
42908 | 274 |
equations: (@{syntax thmdecl}? @{syntax prop} + '|') |
26849 | 275 |
; |
42908 | 276 |
functionopts: '(' (('sequential' | 'domintros') + ',') ')' |
26849 | 277 |
; |
42908 | 278 |
@@{command (HOL) termination} @{syntax term}? |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53982
diff
changeset
|
279 |
; |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53982
diff
changeset
|
280 |
@@{command (HOL) fun_cases} (@{syntax thmdecl}? @{syntax prop} + @'and') |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
281 |
\<close>} |
26849 | 282 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
283 |
\begin{description} |
42123 | 284 |
|
42908 | 285 |
\item @{command (HOL) "primrec"} defines primitive recursive |
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
286 |
functions over datatypes (see also @{command_ref (HOL) datatype_new}). |
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
287 |
The given @{text equations} specify reduction rules that are produced |
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
288 |
by instantiating the generic combinator for primitive recursion that |
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
289 |
is available for each datatype. |
42912 | 290 |
|
291 |
Each equation needs to be of the form: |
|
292 |
||
293 |
@{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"} |
|
294 |
||
295 |
such that @{text C} is a datatype constructor, @{text rhs} contains |
|
296 |
only the free variables on the left-hand side (or from the context), |
|
297 |
and all recursive occurrences of @{text "f"} in @{text "rhs"} are of |
|
298 |
the form @{text "f \<dots> y\<^sub>i \<dots>"} for some @{text i}. At most one |
|
299 |
reduction rule for each constructor can be given. The order does |
|
300 |
not matter. For missing constructors, the function is defined to |
|
301 |
return a default value, but this equation is made difficult to |
|
302 |
access for users. |
|
303 |
||
304 |
The reduction rules are declared as @{attribute simp} by default, |
|
305 |
which enables standard proof methods like @{method simp} and |
|
306 |
@{method auto} to normalize expressions of @{text "f"} applied to |
|
307 |
datatype constructions, by simulating symbolic computation via |
|
308 |
rewriting. |
|
35744 | 309 |
|
42908 | 310 |
\item @{command (HOL) "function"} defines functions by general |
311 |
wellfounded recursion. A detailed description with examples can be |
|
312 |
found in \cite{isabelle-function}. The function is specified by a |
|
313 |
set of (possibly conditional) recursive equations with arbitrary |
|
314 |
pattern matching. The command generates proof obligations for the |
|
315 |
completeness and the compatibility of patterns. |
|
42907
dfd4ef8e73f6
updated and re-unified HOL typedef, with some live examples;
wenzelm
parents:
42705
diff
changeset
|
316 |
|
42908 | 317 |
The defined function is considered partial, and the resulting |
318 |
simplification rules (named @{text "f.psimps"}) and induction rule |
|
319 |
(named @{text "f.pinduct"}) are guarded by a generated domain |
|
320 |
predicate @{text "f_dom"}. The @{command (HOL) "termination"} |
|
321 |
command can then be used to establish that the function is total. |
|
42123 | 322 |
|
42908 | 323 |
\item @{command (HOL) "fun"} is a shorthand notation for ``@{command |
324 |
(HOL) "function"}~@{text "(sequential)"}, followed by automated |
|
325 |
proof attempts regarding pattern matching and termination. See |
|
326 |
\cite{isabelle-function} for further details. |
|
42123 | 327 |
|
42908 | 328 |
\item @{command (HOL) "termination"}~@{text f} commences a |
329 |
termination proof for the previously defined function @{text f}. If |
|
330 |
this is omitted, the command refers to the most recent function |
|
331 |
definition. After the proof is closed, the recursive equations and |
|
332 |
the induction principle is established. |
|
26849 | 333 |
|
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53982
diff
changeset
|
334 |
\item @{command (HOL) "fun_cases"} generates specialized elimination |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53982
diff
changeset
|
335 |
rules for function equations. It expects one or more function equations |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53982
diff
changeset
|
336 |
and produces rules that eliminate the given equalities, following the cases |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53982
diff
changeset
|
337 |
given in the function definition. |
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
338 |
\end{description} |
42907
dfd4ef8e73f6
updated and re-unified HOL typedef, with some live examples;
wenzelm
parents:
42705
diff
changeset
|
339 |
|
42908 | 340 |
Recursive definitions introduced by the @{command (HOL) "function"} |
42912 | 341 |
command accommodate reasoning by induction (cf.\ @{method induct}): |
342 |
rule @{text "f.induct"} refers to a specific induction rule, with |
|
343 |
parameters named according to the user-specified equations. Cases |
|
344 |
are numbered starting from 1. For @{command (HOL) "primrec"}, the |
|
345 |
induction principle coincides with structural recursion on the |
|
346 |
datatype where the recursion is carried out. |
|
42908 | 347 |
|
348 |
The equations provided by these packages may be referred later as |
|
349 |
theorem list @{text "f.simps"}, where @{text f} is the (collective) |
|
350 |
name of the functions defined. Individual equations may be named |
|
351 |
explicitly as well. |
|
352 |
||
353 |
The @{command (HOL) "function"} command accepts the following |
|
354 |
options. |
|
355 |
||
356 |
\begin{description} |
|
357 |
||
358 |
\item @{text sequential} enables a preprocessor which disambiguates |
|
359 |
overlapping patterns by making them mutually disjoint. Earlier |
|
360 |
equations take precedence over later ones. This allows to give the |
|
361 |
specification in a format very similar to functional programming. |
|
362 |
Note that the resulting simplification and induction rules |
|
363 |
correspond to the transformed specification, not the one given |
|
364 |
originally. This usually means that each equation given by the user |
|
365 |
may result in several theorems. Also note that this automatic |
|
366 |
transformation only works for ML-style datatype patterns. |
|
367 |
||
368 |
\item @{text domintros} enables the automated generation of |
|
369 |
introduction rules for the domain predicate. While mostly not |
|
370 |
needed, they can be helpful in some proofs about partial functions. |
|
371 |
||
372 |
\end{description} |
|
26849 | 373 |
*} |
374 |
||
42912 | 375 |
subsubsection {* Example: evaluation of expressions *} |
376 |
||
377 |
text {* Subsequently, we define mutual datatypes for arithmetic and |
|
378 |
boolean expressions, and use @{command primrec} for evaluation |
|
379 |
functions that follow the same recursive structure. *} |
|
380 |
||
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
381 |
datatype_new 'a aexp = |
42912 | 382 |
IF "'a bexp" "'a aexp" "'a aexp" |
383 |
| Sum "'a aexp" "'a aexp" |
|
384 |
| Diff "'a aexp" "'a aexp" |
|
385 |
| Var 'a |
|
386 |
| Num nat |
|
387 |
and 'a bexp = |
|
388 |
Less "'a aexp" "'a aexp" |
|
389 |
| And "'a bexp" "'a bexp" |
|
390 |
| Neg "'a bexp" |
|
391 |
||
392 |
||
393 |
text {* \medskip Evaluation of arithmetic and boolean expressions *} |
|
394 |
||
395 |
primrec evala :: "('a \<Rightarrow> nat) \<Rightarrow> 'a aexp \<Rightarrow> nat" |
|
396 |
and evalb :: "('a \<Rightarrow> nat) \<Rightarrow> 'a bexp \<Rightarrow> bool" |
|
397 |
where |
|
398 |
"evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)" |
|
399 |
| "evala env (Sum a1 a2) = evala env a1 + evala env a2" |
|
400 |
| "evala env (Diff a1 a2) = evala env a1 - evala env a2" |
|
401 |
| "evala env (Var v) = env v" |
|
402 |
| "evala env (Num n) = n" |
|
403 |
| "evalb env (Less a1 a2) = (evala env a1 < evala env a2)" |
|
404 |
| "evalb env (And b1 b2) = (evalb env b1 \<and> evalb env b2)" |
|
405 |
| "evalb env (Neg b) = (\<not> evalb env b)" |
|
406 |
||
407 |
text {* Since the value of an expression depends on the value of its |
|
408 |
variables, the functions @{const evala} and @{const evalb} take an |
|
409 |
additional parameter, an \emph{environment} that maps variables to |
|
410 |
their values. |
|
411 |
||
412 |
\medskip Substitution on expressions can be defined similarly. The |
|
413 |
mapping @{text f} of type @{typ "'a \<Rightarrow> 'a aexp"} given as a |
|
414 |
parameter is lifted canonically on the types @{typ "'a aexp"} and |
|
415 |
@{typ "'a bexp"}, respectively. |
|
416 |
*} |
|
417 |
||
418 |
primrec substa :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a aexp \<Rightarrow> 'b aexp" |
|
419 |
and substb :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a bexp \<Rightarrow> 'b bexp" |
|
420 |
where |
|
421 |
"substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)" |
|
422 |
| "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)" |
|
423 |
| "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)" |
|
424 |
| "substa f (Var v) = f v" |
|
425 |
| "substa f (Num n) = Num n" |
|
426 |
| "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)" |
|
427 |
| "substb f (And b1 b2) = And (substb f b1) (substb f b2)" |
|
428 |
| "substb f (Neg b) = Neg (substb f b)" |
|
429 |
||
430 |
text {* In textbooks about semantics one often finds substitution |
|
431 |
theorems, which express the relationship between substitution and |
|
432 |
evaluation. For @{typ "'a aexp"} and @{typ "'a bexp"}, we can prove |
|
433 |
such a theorem by mutual induction, followed by simplification. |
|
434 |
*} |
|
435 |
||
436 |
lemma subst_one: |
|
437 |
"evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a" |
|
438 |
"evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b" |
|
439 |
by (induct a and b) simp_all |
|
440 |
||
441 |
lemma subst_all: |
|
442 |
"evala env (substa s a) = evala (\<lambda>x. evala env (s x)) a" |
|
443 |
"evalb env (substb s b) = evalb (\<lambda>x. evala env (s x)) b" |
|
444 |
by (induct a and b) simp_all |
|
445 |
||
446 |
||
447 |
subsubsection {* Example: a substitution function for terms *} |
|
448 |
||
449 |
text {* Functions on datatypes with nested recursion are also defined |
|
450 |
by mutual primitive recursion. *} |
|
451 |
||
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
452 |
datatype_new ('a, 'b) "term" = Var 'a | App 'b "('a, 'b) term list" |
42912 | 453 |
|
454 |
text {* A substitution function on type @{typ "('a, 'b) term"} can be |
|
455 |
defined as follows, by working simultaneously on @{typ "('a, 'b) |
|
456 |
term list"}: *} |
|
457 |
||
458 |
primrec subst_term :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term \<Rightarrow> ('a, 'b) term" and |
|
459 |
subst_term_list :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term list \<Rightarrow> ('a, 'b) term list" |
|
460 |
where |
|
461 |
"subst_term f (Var a) = f a" |
|
462 |
| "subst_term f (App b ts) = App b (subst_term_list f ts)" |
|
463 |
| "subst_term_list f [] = []" |
|
464 |
| "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts" |
|
465 |
||
466 |
text {* The recursion scheme follows the structure of the unfolded |
|
467 |
definition of type @{typ "('a, 'b) term"}. To prove properties of this |
|
468 |
substitution function, mutual induction is needed: |
|
469 |
*} |
|
470 |
||
471 |
lemma "subst_term (subst_term f1 \<circ> f2) t = subst_term f1 (subst_term f2 t)" and |
|
472 |
"subst_term_list (subst_term f1 \<circ> f2) ts = subst_term_list f1 (subst_term_list f2 ts)" |
|
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
473 |
by (induct t and ts rule: subst_term.induct subst_term_list.induct) simp_all |
42912 | 474 |
|
475 |
||
476 |
subsubsection {* Example: a map function for infinitely branching trees *} |
|
477 |
||
478 |
text {* Defining functions on infinitely branching datatypes by |
|
479 |
primitive recursion is just as easy. |
|
480 |
*} |
|
481 |
||
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
482 |
datatype_new 'a tree = Atom 'a | Branch "nat \<Rightarrow> 'a tree" |
42912 | 483 |
|
484 |
primrec map_tree :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a tree \<Rightarrow> 'b tree" |
|
485 |
where |
|
486 |
"map_tree f (Atom a) = Atom (f a)" |
|
487 |
| "map_tree f (Branch ts) = Branch (\<lambda>x. map_tree f (ts x))" |
|
488 |
||
489 |
text {* Note that all occurrences of functions such as @{text ts} |
|
490 |
above must be applied to an argument. In particular, @{term |
|
491 |
"map_tree f \<circ> ts"} is not allowed here. *} |
|
492 |
||
493 |
text {* Here is a simple composition lemma for @{term map_tree}: *} |
|
494 |
||
495 |
lemma "map_tree g (map_tree f t) = map_tree (g \<circ> f) t" |
|
496 |
by (induct t) simp_all |
|
497 |
||
42907
dfd4ef8e73f6
updated and re-unified HOL typedef, with some live examples;
wenzelm
parents:
42705
diff
changeset
|
498 |
|
42908 | 499 |
subsection {* Proof methods related to recursive definitions *} |
26849 | 500 |
|
501 |
text {* |
|
502 |
\begin{matharray}{rcl} |
|
42908 | 503 |
@{method_def (HOL) pat_completeness} & : & @{text method} \\ |
504 |
@{method_def (HOL) relation} & : & @{text method} \\ |
|
505 |
@{method_def (HOL) lexicographic_order} & : & @{text method} \\ |
|
506 |
@{method_def (HOL) size_change} & : & @{text method} \\ |
|
45944
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
bulwahn
parents:
45943
diff
changeset
|
507 |
@{method_def (HOL) induction_schema} & : & @{text method} \\ |
26849 | 508 |
\end{matharray} |
509 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
510 |
@{rail \<open> |
42908 | 511 |
@@{method (HOL) relation} @{syntax term} |
512 |
; |
|
513 |
@@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * ) |
|
514 |
; |
|
515 |
@@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) ) |
|
516 |
; |
|
45944
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
bulwahn
parents:
45943
diff
changeset
|
517 |
@@{method (HOL) induction_schema} |
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
bulwahn
parents:
45943
diff
changeset
|
518 |
; |
42908 | 519 |
orders: ( 'max' | 'min' | 'ms' ) * |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
520 |
\<close>} |
42908 | 521 |
|
522 |
\begin{description} |
|
523 |
||
524 |
\item @{method (HOL) pat_completeness} is a specialized method to |
|
525 |
solve goals regarding the completeness of pattern matching, as |
|
526 |
required by the @{command (HOL) "function"} package (cf.\ |
|
527 |
\cite{isabelle-function}). |
|
528 |
||
529 |
\item @{method (HOL) relation}~@{text R} introduces a termination |
|
530 |
proof using the relation @{text R}. The resulting proof state will |
|
531 |
contain goals expressing that @{text R} is wellfounded, and that the |
|
532 |
arguments of recursive calls decrease with respect to @{text R}. |
|
533 |
Usually, this method is used as the initial proof step of manual |
|
534 |
termination proofs. |
|
535 |
||
536 |
\item @{method (HOL) "lexicographic_order"} attempts a fully |
|
537 |
automated termination proof by searching for a lexicographic |
|
538 |
combination of size measures on the arguments of the function. The |
|
539 |
method accepts the same arguments as the @{method auto} method, |
|
42930 | 540 |
which it uses internally to prove local descents. The @{syntax |
541 |
clasimpmod} modifiers are accepted (as for @{method auto}). |
|
42908 | 542 |
|
543 |
In case of failure, extensive information is printed, which can help |
|
544 |
to analyse the situation (cf.\ \cite{isabelle-function}). |
|
545 |
||
546 |
\item @{method (HOL) "size_change"} also works on termination goals, |
|
547 |
using a variation of the size-change principle, together with a |
|
548 |
graph decomposition technique (see \cite{krauss_phd} for details). |
|
549 |
Three kinds of orders are used internally: @{text max}, @{text min}, |
|
550 |
and @{text ms} (multiset), which is only available when the theory |
|
551 |
@{text Multiset} is loaded. When no order kinds are given, they are |
|
552 |
tried in order. The search for a termination proof uses SAT solving |
|
553 |
internally. |
|
554 |
||
42930 | 555 |
For local descent proofs, the @{syntax clasimpmod} modifiers are |
556 |
accepted (as for @{method auto}). |
|
42908 | 557 |
|
45944
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
bulwahn
parents:
45943
diff
changeset
|
558 |
\item @{method (HOL) induction_schema} derives user-specified |
46283 | 559 |
induction rules from well-founded induction and completeness of |
560 |
patterns. This factors out some operations that are done internally |
|
561 |
by the function package and makes them available separately. See |
|
562 |
@{file "~~/src/HOL/ex/Induction_Schema.thy"} for examples. |
|
45944
e586f6d136b7
added some basic documentation about method induction_schema extracted from old NEWS
bulwahn
parents:
45943
diff
changeset
|
563 |
|
42908 | 564 |
\end{description} |
565 |
*} |
|
566 |
||
567 |
||
568 |
subsection {* Functions with explicit partiality *} |
|
569 |
||
570 |
text {* |
|
571 |
\begin{matharray}{rcl} |
|
572 |
@{command_def (HOL) "partial_function"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
|
573 |
@{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\ |
|
574 |
\end{matharray} |
|
575 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
576 |
@{rail \<open> |
42908 | 577 |
@@{command (HOL) partial_function} @{syntax target}? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
578 |
'(' @{syntax nameref} ')' @{syntax "fixes"} \<newline> |
42908 | 579 |
@'where' @{syntax thmdecl}? @{syntax prop} |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
580 |
\<close>} |
26849 | 581 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
582 |
\begin{description} |
42123 | 583 |
|
42908 | 584 |
\item @{command (HOL) "partial_function"}~@{text "(mode)"} defines |
585 |
recursive functions based on fixpoints in complete partial |
|
586 |
orders. No termination proof is required from the user or |
|
587 |
constructed internally. Instead, the possibility of non-termination |
|
588 |
is modelled explicitly in the result type, which contains an |
|
589 |
explicit bottom element. |
|
590 |
||
591 |
Pattern matching and mutual recursion are currently not supported. |
|
592 |
Thus, the specification consists of a single function described by a |
|
593 |
single recursive equation. |
|
594 |
||
595 |
There are no fixed syntactic restrictions on the body of the |
|
596 |
function, but the induced functional must be provably monotonic |
|
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
597 |
wrt.\ the underlying order. The monotonicity proof is performed |
42908 | 598 |
internally, and the definition is rejected when it fails. The proof |
599 |
can be influenced by declaring hints using the |
|
600 |
@{attribute (HOL) partial_function_mono} attribute. |
|
601 |
||
602 |
The mandatory @{text mode} argument specifies the mode of operation |
|
603 |
of the command, which directly corresponds to a complete partial |
|
604 |
order on the result type. By default, the following modes are |
|
605 |
defined: |
|
26849 | 606 |
|
42908 | 607 |
\begin{description} |
46283 | 608 |
|
42908 | 609 |
\item @{text option} defines functions that map into the @{type |
610 |
option} type. Here, the value @{term None} is used to model a |
|
611 |
non-terminating computation. Monotonicity requires that if @{term |
|
46283 | 612 |
None} is returned by a recursive call, then the overall result must |
613 |
also be @{term None}. This is best achieved through the use of the |
|
614 |
monadic operator @{const "Option.bind"}. |
|
42908 | 615 |
|
616 |
\item @{text tailrec} defines functions with an arbitrary result |
|
617 |
type and uses the slightly degenerated partial order where @{term |
|
618 |
"undefined"} is the bottom element. Now, monotonicity requires that |
|
619 |
if @{term undefined} is returned by a recursive call, then the |
|
620 |
overall result must also be @{term undefined}. In practice, this is |
|
621 |
only satisfied when each recursive call is a tail call, whose result |
|
622 |
is directly returned. Thus, this mode of operation allows the |
|
623 |
definition of arbitrary tail-recursive functions. |
|
46283 | 624 |
|
42908 | 625 |
\end{description} |
626 |
||
627 |
Experienced users may define new modes by instantiating the locale |
|
628 |
@{const "partial_function_definitions"} appropriately. |
|
629 |
||
630 |
\item @{attribute (HOL) partial_function_mono} declares rules for |
|
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
631 |
use in the internal monotonicity proofs of partial function |
42908 | 632 |
definitions. |
26849 | 633 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
634 |
\end{description} |
42908 | 635 |
|
636 |
*} |
|
637 |
||
638 |
||
639 |
subsection {* Old-style recursive function definitions (TFL) *} |
|
640 |
||
641 |
text {* |
|
642 |
\begin{matharray}{rcl} |
|
643 |
@{command_def (HOL) "recdef"} & : & @{text "theory \<rightarrow> theory)"} \\ |
|
644 |
@{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\ |
|
645 |
\end{matharray} |
|
646 |
||
46280 | 647 |
The old TFL commands @{command (HOL) "recdef"} and @{command (HOL) |
648 |
"recdef_tc"} for defining recursive are mostly obsolete; @{command |
|
649 |
(HOL) "function"} or @{command (HOL) "fun"} should be used instead. |
|
650 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
651 |
@{rail \<open> |
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
652 |
@@{command (HOL) recdef} ('(' @'permissive' ')')? \<newline> |
42908 | 653 |
@{syntax name} @{syntax term} (@{syntax prop} +) hints? |
654 |
; |
|
655 |
recdeftc @{syntax thmdecl}? tc |
|
656 |
; |
|
657 |
hints: '(' @'hints' ( recdefmod * ) ')' |
|
658 |
; |
|
659 |
recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf') |
|
660 |
(() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod} |
|
661 |
; |
|
662 |
tc: @{syntax nameref} ('(' @{syntax nat} ')')? |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
663 |
\<close>} |
42908 | 664 |
|
665 |
\begin{description} |
|
666 |
||
667 |
\item @{command (HOL) "recdef"} defines general well-founded |
|
668 |
recursive functions (using the TFL package), see also |
|
669 |
\cite{isabelle-HOL}. The ``@{text "(permissive)"}'' option tells |
|
670 |
TFL to recover from failed proof attempts, returning unfinished |
|
671 |
results. The @{text recdef_simp}, @{text recdef_cong}, and @{text |
|
672 |
recdef_wf} hints refer to auxiliary rules to be used in the internal |
|
673 |
automated proof process of TFL. Additional @{syntax clasimpmod} |
|
42930 | 674 |
declarations may be given to tune the context of the Simplifier |
675 |
(cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\ |
|
676 |
\secref{sec:classical}). |
|
42908 | 677 |
|
678 |
\item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the |
|
679 |
proof for leftover termination condition number @{text i} (default |
|
680 |
1) as generated by a @{command (HOL) "recdef"} definition of |
|
681 |
constant @{text c}. |
|
682 |
||
683 |
Note that in most cases, @{command (HOL) "recdef"} is able to finish |
|
684 |
its internal proofs without manual intervention. |
|
685 |
||
686 |
\end{description} |
|
687 |
||
688 |
\medskip Hints for @{command (HOL) "recdef"} may be also declared |
|
689 |
globally, using the following attributes. |
|
690 |
||
691 |
\begin{matharray}{rcl} |
|
692 |
@{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\ |
|
693 |
@{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\ |
|
694 |
@{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\ |
|
695 |
\end{matharray} |
|
696 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
697 |
@{rail \<open> |
42908 | 698 |
(@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} | |
699 |
@@{attribute (HOL) recdef_wf}) (() | 'add' | 'del') |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
700 |
\<close>} |
42908 | 701 |
*} |
702 |
||
703 |
||
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
704 |
section {* Old-style datatypes \label{sec:hol-datatype} *} |
42908 | 705 |
|
706 |
text {* |
|
707 |
\begin{matharray}{rcl} |
|
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
708 |
@{command_def (HOL) "old_datatype"} & : & @{text "theory \<rightarrow> theory"} \\ |
58306
117ba6cbe414
renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
blanchet
parents:
58305
diff
changeset
|
709 |
@{command_def (HOL) "old_rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\ |
42908 | 710 |
\end{matharray} |
711 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
712 |
@{rail \<open> |
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
713 |
@@{command (HOL) old_datatype} (spec + @'and') |
42908 | 714 |
; |
58306
117ba6cbe414
renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
blanchet
parents:
58305
diff
changeset
|
715 |
@@{command (HOL) old_rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +) |
42908 | 716 |
; |
717 |
||
45839
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
wenzelm
parents:
45768
diff
changeset
|
718 |
spec: @{syntax typespec_sorts} @{syntax mixfix}? '=' (cons + '|') |
42908 | 719 |
; |
720 |
cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}? |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
721 |
\<close>} |
42908 | 722 |
|
723 |
\begin{description} |
|
724 |
||
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
725 |
\item @{command (HOL) "old_datatype"} defines old-style inductive |
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
726 |
datatypes in HOL. |
42908 | 727 |
|
58306
117ba6cbe414
renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
blanchet
parents:
58305
diff
changeset
|
728 |
\item @{command (HOL) "old_rep_datatype"} represents existing types as |
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
729 |
old-style datatypes. |
42908 | 730 |
|
731 |
\end{description} |
|
732 |
||
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
733 |
These commands are mostly obsolete; @{command (HOL) "datatype"} |
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
734 |
should be used instead. |
42908 | 735 |
|
736 |
See \cite{isabelle-HOL} for more details on datatypes, but beware of |
|
737 |
the old-style theory syntax being used there! Apart from proper |
|
738 |
proof methods for case-analysis and induction, there are also |
|
739 |
emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL) |
|
740 |
induct_tac} available, see \secref{sec:hol-induct-tac}; these admit |
|
741 |
to refer directly to the internal structure of subgoals (including |
|
742 |
internally bound parameters). |
|
26849 | 743 |
*} |
744 |
||
745 |
||
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
|
746 |
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
|
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 |
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
|
749 |
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
|
750 |
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
|
751 |
|
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
752 |
datatype_new 'a seq = Empty | Seq 'a "'a seq" |
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
|
753 |
|
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 |
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
|
755 |
|
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 |
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
|
757 |
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
|
758 |
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
|
759 |
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
|
760 |
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
|
761 |
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
|
762 |
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
|
763 |
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
|
764 |
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
|
765 |
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
|
766 |
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
|
767 |
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
|
768 |
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
|
769 |
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
|
770 |
|
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
|
771 |
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
|
772 |
|
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
|
773 |
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
|
774 |
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
|
775 |
|
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
|
776 |
|
26849 | 777 |
section {* Records \label{sec:hol-record} *} |
778 |
||
779 |
text {* |
|
780 |
In principle, records merely generalize the concept of tuples, where |
|
781 |
components may be addressed by labels instead of just position. The |
|
782 |
logical infrastructure of records in Isabelle/HOL is slightly more |
|
783 |
advanced, though, supporting truly extensible record schemes. This |
|
784 |
admits operations that are polymorphic with respect to record |
|
785 |
extension, yielding ``object-oriented'' effects like (single) |
|
786 |
inheritance. See also \cite{NaraschewskiW-TPHOLs98} for more |
|
787 |
details on object-oriented verification and record subtyping in HOL. |
|
788 |
*} |
|
789 |
||
790 |
||
791 |
subsection {* Basic concepts *} |
|
792 |
||
793 |
text {* |
|
794 |
Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records |
|
795 |
at the level of terms and types. The notation is as follows: |
|
796 |
||
797 |
\begin{center} |
|
798 |
\begin{tabular}{l|l|l} |
|
799 |
& record terms & record types \\ \hline |
|
800 |
fixed & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\ |
|
801 |
schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} & |
|
802 |
@{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\ |
|
803 |
\end{tabular} |
|
804 |
\end{center} |
|
805 |
||
806 |
\noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text |
|
807 |
"(| x = a |)"}. |
|
808 |
||
809 |
A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value |
|
810 |
@{text a} and field @{text y} of value @{text b}. The corresponding |
|
811 |
type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"} |
|
812 |
and @{text "b :: B"}. |
|
813 |
||
814 |
A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields |
|
815 |
@{text x} and @{text y} as before, but also possibly further fields |
|
816 |
as indicated by the ``@{text "\<dots>"}'' notation (which is actually part |
|
817 |
of the syntax). The improper field ``@{text "\<dots>"}'' of a record |
|
818 |
scheme is called the \emph{more part}. Logically it is just a free |
|
819 |
variable, which is occasionally referred to as ``row variable'' in |
|
820 |
the literature. The more part of a record scheme may be |
|
821 |
instantiated by zero or more further components. For example, the |
|
822 |
previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z = |
|
26852 | 823 |
c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part. |
26849 | 824 |
Fixed records are special instances of record schemes, where |
825 |
``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"} |
|
826 |
element. In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation |
|
827 |
for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}. |
|
42123 | 828 |
|
26849 | 829 |
\medskip Two key observations make extensible records in a simply |
830 |
typed language like HOL work out: |
|
831 |
||
832 |
\begin{enumerate} |
|
833 |
||
834 |
\item the more part is internalized, as a free term or type |
|
835 |
variable, |
|
836 |
||
26852 | 837 |
\item field names are externalized, they cannot be accessed within |
838 |
the logic as first-class values. |
|
26849 | 839 |
|
840 |
\end{enumerate} |
|
841 |
||
842 |
\medskip In Isabelle/HOL record types have to be defined explicitly, |
|
843 |
fixing their field names and types, and their (optional) parent |
|
844 |
record. Afterwards, records may be formed using above syntax, while |
|
845 |
obeying the canonical order of fields as given by their declaration. |
|
846 |
The record package provides several standard operations like |
|
847 |
selectors and updates. The common setup for various generic proof |
|
848 |
tools enable succinct reasoning patterns. See also the Isabelle/HOL |
|
849 |
tutorial \cite{isabelle-hol-book} for further instructions on using |
|
850 |
records in practice. |
|
851 |
*} |
|
852 |
||
853 |
||
854 |
subsection {* Record specifications *} |
|
855 |
||
856 |
text {* |
|
857 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
858 |
@{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\ |
26849 | 859 |
\end{matharray} |
860 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
861 |
@{rail \<open> |
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
862 |
@@{command (HOL) record} @{syntax typespec_sorts} '=' \<newline> |
46494
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
wenzelm
parents:
46457
diff
changeset
|
863 |
(@{syntax type} '+')? (constdecl +) |
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
wenzelm
parents:
46457
diff
changeset
|
864 |
; |
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
wenzelm
parents:
46457
diff
changeset
|
865 |
constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
866 |
\<close>} |
26849 | 867 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
868 |
\begin{description} |
26849 | 869 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
870 |
\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
|
871 |
\<dots> c\<^sub>n :: \<sigma>\<^sub>n"} defines extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"}, |
26849 | 872 |
derived from the optional parent record @{text "\<tau>"} by adding new |
873 |
field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc. |
|
874 |
||
875 |
The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be |
|
876 |
covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>, |
|
877 |
\<alpha>\<^sub>m"}. Type constructor @{text t} has to be new, while @{text |
|
878 |
\<tau>} needs to specify an instance of an existing record type. At |
|
879 |
least one new field @{text "c\<^sub>i"} has to be specified. |
|
880 |
Basically, field names need to belong to a unique record. This is |
|
881 |
not a real restriction in practice, since fields are qualified by |
|
882 |
the record name internally. |
|
883 |
||
884 |
The parent record specification @{text \<tau>} is optional; if omitted |
|
885 |
@{text t} becomes a root record. The hierarchy of all records |
|
886 |
declared within a theory context forms a forest structure, i.e.\ a |
|
887 |
set of trees starting with a root record each. There is no way to |
|
888 |
merge multiple parent records! |
|
889 |
||
890 |
For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a |
|
891 |
type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 :: |
|
892 |
\<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text |
|
893 |
"(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for |
|
894 |
@{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> :: |
|
895 |
\<zeta>\<rparr>"}. |
|
896 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
897 |
\end{description} |
26849 | 898 |
*} |
899 |
||
900 |
||
901 |
subsection {* Record operations *} |
|
902 |
||
903 |
text {* |
|
904 |
Any record definition of the form presented above produces certain |
|
905 |
standard operations. Selectors and updates are provided for any |
|
906 |
field, including the improper one ``@{text more}''. There are also |
|
907 |
cumulative record constructor functions. To simplify the |
|
908 |
presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>, |
|
909 |
\<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 :: |
|
910 |
\<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}. |
|
911 |
||
912 |
\medskip \textbf{Selectors} and \textbf{updates} are available for |
|
913 |
any field (including ``@{text more}''): |
|
914 |
||
915 |
\begin{matharray}{lll} |
|
26852 | 916 |
@{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\ |
917 |
@{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 | 918 |
\end{matharray} |
919 |
||
920 |
There is special syntax for application of updates: @{text "r\<lparr>x := |
|
921 |
a\<rparr>"} abbreviates term @{text "x_update a r"}. Further notation for |
|
922 |
repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z := |
|
923 |
c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}. Note that |
|
924 |
because of postfix notation the order of fields shown here is |
|
925 |
reverse than in the actual term. Since repeated updates are just |
|
926 |
function applications, fields may be freely permuted in @{text "\<lparr>x |
|
927 |
:= a, y := b, z := c\<rparr>"}, as far as logical equality is concerned. |
|
928 |
Thus commutativity of independent updates can be proven within the |
|
929 |
logic for any two fields, but not as a general theorem. |
|
930 |
||
931 |
\medskip The \textbf{make} operation provides a cumulative record |
|
932 |
constructor function: |
|
933 |
||
934 |
\begin{matharray}{lll} |
|
26852 | 935 |
@{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\ |
26849 | 936 |
\end{matharray} |
937 |
||
938 |
\medskip We now reconsider the case of non-root records, which are |
|
939 |
derived of some parent. In general, the latter may depend on |
|
940 |
another parent as well, resulting in a list of \emph{ancestor |
|
941 |
records}. Appending the lists of fields of all ancestors results in |
|
942 |
a certain field prefix. The record package automatically takes care |
|
943 |
of this by lifting operations over this context of ancestor fields. |
|
944 |
Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor |
|
945 |
fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"}, |
|
946 |
the above record operations will get the following types: |
|
947 |
||
26852 | 948 |
\medskip |
949 |
\begin{tabular}{lll} |
|
950 |
@{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\ |
|
42123 | 951 |
@{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow> |
26852 | 952 |
\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> |
953 |
\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\ |
|
954 |
@{text "t.make"} & @{text "::"} & @{text "\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> |
|
955 |
\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\ |
|
956 |
\end{tabular} |
|
957 |
\medskip |
|
26849 | 958 |
|
26852 | 959 |
\noindent Some further operations address the extension aspect of a |
26849 | 960 |
derived record scheme specifically: @{text "t.fields"} produces a |
961 |
record fragment consisting of exactly the new fields introduced here |
|
962 |
(the result may serve as a more part elsewhere); @{text "t.extend"} |
|
963 |
takes a fixed record and adds a given more part; @{text |
|
964 |
"t.truncate"} restricts a record scheme to a fixed record. |
|
965 |
||
26852 | 966 |
\medskip |
967 |
\begin{tabular}{lll} |
|
968 |
@{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\ |
|
969 |
@{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow> |
|
970 |
\<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\ |
|
971 |
@{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>"} \\ |
|
972 |
\end{tabular} |
|
973 |
\medskip |
|
26849 | 974 |
|
975 |
\noindent Note that @{text "t.make"} and @{text "t.fields"} coincide |
|
976 |
for root records. |
|
977 |
*} |
|
978 |
||
979 |
||
980 |
subsection {* Derived rules and proof tools *} |
|
981 |
||
982 |
text {* |
|
983 |
The record package proves several results internally, declaring |
|
984 |
these facts to appropriate proof tools. This enables users to |
|
985 |
reason about record structures quite conveniently. Assume that |
|
986 |
@{text t} is a record type as specified above. |
|
987 |
||
988 |
\begin{enumerate} |
|
42123 | 989 |
|
26849 | 990 |
\item Standard conversions for selectors or updates applied to |
991 |
record constructor terms are made part of the default Simplifier |
|
992 |
context; thus proofs by reduction of basic operations merely require |
|
993 |
the @{method simp} method without further arguments. These rules |
|
994 |
are available as @{text "t.simps"}, too. |
|
42123 | 995 |
|
26849 | 996 |
\item Selectors applied to updated records are automatically reduced |
997 |
by an internal simplification procedure, which is also part of the |
|
998 |
standard Simplifier setup. |
|
999 |
||
1000 |
\item Inject equations of a form analogous to @{prop "(x, y) = (x', |
|
1001 |
y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical |
|
1002 |
Reasoner as @{attribute iff} rules. These rules are available as |
|
1003 |
@{text "t.iffs"}. |
|
1004 |
||
1005 |
\item The introduction rule for record equality analogous to @{text |
|
1006 |
"x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier, |
|
1007 |
and as the basic rule context as ``@{attribute intro}@{text "?"}''. |
|
1008 |
The rule is called @{text "t.equality"}. |
|
1009 |
||
1010 |
\item Representations of arbitrary record expressions as canonical |
|
1011 |
constructor terms are provided both in @{method cases} and @{method |
|
1012 |
induct} format (cf.\ the generic proof methods of the same name, |
|
1013 |
\secref{sec:cases-induct}). Several variations are available, for |
|
1014 |
fixed records, record schemes, more parts etc. |
|
42123 | 1015 |
|
26849 | 1016 |
The generic proof methods are sufficiently smart to pick the most |
1017 |
sensible rule according to the type of the indicated record |
|
1018 |
expression: users just need to apply something like ``@{text "(cases |
|
1019 |
r)"}'' to a certain proof problem. |
|
1020 |
||
1021 |
\item The derived record operations @{text "t.make"}, @{text |
|
1022 |
"t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not} |
|
1023 |
treated automatically, but usually need to be expanded by hand, |
|
1024 |
using the collective fact @{text "t.defs"}. |
|
1025 |
||
1026 |
\end{enumerate} |
|
1027 |
*} |
|
1028 |
||
1029 |
||
42911 | 1030 |
subsubsection {* Examples *} |
1031 |
||
1032 |
text {* See @{file "~~/src/HOL/ex/Records.thy"}, for example. *} |
|
1033 |
||
42908 | 1034 |
section {* Typedef axiomatization \label{sec:hol-typedef} *} |
1035 |
||
46280 | 1036 |
text {* |
1037 |
\begin{matharray}{rcl} |
|
1038 |
@{command_def (HOL) "typedef"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\ |
|
1039 |
\end{matharray} |
|
1040 |
||
1041 |
A Gordon/HOL-style type definition is a certain axiom scheme that |
|
1042 |
identifies a new type with a subset of an existing type. More |
|
42908 | 1043 |
precisely, the new type is defined by exhibiting an existing type |
1044 |
@{text \<tau>}, a set @{text "A :: \<tau> set"}, and a theorem that proves |
|
1045 |
@{prop "\<exists>x. x \<in> A"}. Thus @{text A} is a non-empty subset of @{text |
|
1046 |
\<tau>}, and the new type denotes this subset. New functions are |
|
1047 |
postulated that establish an isomorphism between the new type and |
|
1048 |
the subset. In general, the type @{text \<tau>} may involve type |
|
1049 |
variables @{text "\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n"} which means that the type definition |
|
1050 |
produces a type constructor @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} depending on |
|
1051 |
those type arguments. |
|
1052 |
||
57480
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1053 |
The axiomatization can be considered a ``definition'' in the sense of the |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1054 |
particular set-theoretic interpretation of HOL \cite{pitts93}, where the |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1055 |
universe of types is required to be downwards-closed wrt.\ arbitrary |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1056 |
non-empty subsets. Thus genuinely new types introduced by @{command |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1057 |
"typedef"} stay within the range of HOL models by construction. |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1058 |
|
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1059 |
In contrast, the command @{command_ref type_synonym} from Isabelle/Pure |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1060 |
merely introduces syntactic abbreviations, without any logical |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1061 |
significance. Thus it is more faithful to the idea of a genuine type |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1062 |
definition, but less powerful in practice. |
47859 | 1063 |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1064 |
@{rail \<open> |
49836
c13b39542972
simplified 'typedef' specifications: discontinued implicit set definition and alternative name;
wenzelm
parents:
49834
diff
changeset
|
1065 |
@@{command (HOL) typedef} abs_type '=' rep_set |
26849 | 1066 |
; |
42908 | 1067 |
abs_type: @{syntax typespec_sorts} @{syntax mixfix}? |
1068 |
; |
|
1069 |
rep_set: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})? |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1070 |
\<close>} |
26849 | 1071 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
1072 |
\begin{description} |
26849 | 1073 |
|
57480
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1074 |
\item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"} produces an |
57487 | 1075 |
axiomatization (\secref{sec:axiomatizations}) for a type definition in the |
57480
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1076 |
background theory of the current context, depending on a non-emptiness |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1077 |
result of the set @{text A} that needs to be proven here. The set @{text |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1078 |
A} may contain type variables @{text "\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n"} as specified on the |
d256f49b4799
clarified "axiomatization" -- minor rewording of this delicate concept;
wenzelm
parents:
56929
diff
changeset
|
1079 |
LHS, but no term variables. |
42908 | 1080 |
|
1081 |
Even though a local theory specification, the newly introduced type |
|
1082 |
constructor cannot depend on parameters or assumptions of the |
|
1083 |
context: this is structurally impossible in HOL. In contrast, the |
|
1084 |
non-emptiness proof may use local assumptions in unusual situations, |
|
1085 |
which could result in different interpretations in target contexts: |
|
1086 |
the meaning of the bijection between the representing set @{text A} |
|
1087 |
and the new type @{text t} may then change in different application |
|
1088 |
contexts. |
|
1089 |
||
49836
c13b39542972
simplified 'typedef' specifications: discontinued implicit set definition and alternative name;
wenzelm
parents:
49834
diff
changeset
|
1090 |
For @{command (HOL) "typedef"}~@{text "t = A"} the newly introduced |
c13b39542972
simplified 'typedef' specifications: discontinued implicit set definition and alternative name;
wenzelm
parents:
49834
diff
changeset
|
1091 |
type @{text t} is accompanied by a pair of morphisms to relate it to |
c13b39542972
simplified 'typedef' specifications: discontinued implicit set definition and alternative name;
wenzelm
parents:
49834
diff
changeset
|
1092 |
the representing set over the old type. By default, the injection |
c13b39542972
simplified 'typedef' specifications: discontinued implicit set definition and alternative name;
wenzelm
parents:
49834
diff
changeset
|
1093 |
from type to set is called @{text Rep_t} and its inverse @{text |
c13b39542972
simplified 'typedef' specifications: discontinued implicit set definition and alternative name;
wenzelm
parents:
49834
diff
changeset
|
1094 |
Abs_t}: An explicit @{keyword (HOL) "morphisms"} specification |
c13b39542972
simplified 'typedef' specifications: discontinued implicit set definition and alternative name;
wenzelm
parents:
49834
diff
changeset
|
1095 |
allows to provide alternative names. |
26849 | 1096 |
|
42908 | 1097 |
The core axiomatization uses the locale predicate @{const |
1098 |
type_definition} as defined in Isabelle/HOL. Various basic |
|
1099 |
consequences of that are instantiated accordingly, re-using the |
|
1100 |
locale facts with names derived from the new type constructor. Thus |
|
1101 |
the generic @{thm type_definition.Rep} is turned into the specific |
|
1102 |
@{text "Rep_t"}, for example. |
|
1103 |
||
1104 |
Theorems @{thm type_definition.Rep}, @{thm |
|
1105 |
type_definition.Rep_inverse}, and @{thm type_definition.Abs_inverse} |
|
1106 |
provide the most basic characterization as a corresponding |
|
1107 |
injection/surjection pair (in both directions). The derived rules |
|
1108 |
@{thm type_definition.Rep_inject} and @{thm |
|
1109 |
type_definition.Abs_inject} provide a more convenient version of |
|
1110 |
injectivity, suitable for automated proof tools (e.g.\ in |
|
1111 |
declarations involving @{attribute simp} or @{attribute iff}). |
|
1112 |
Furthermore, the rules @{thm type_definition.Rep_cases}~/ @{thm |
|
1113 |
type_definition.Rep_induct}, and @{thm type_definition.Abs_cases}~/ |
|
1114 |
@{thm type_definition.Abs_induct} provide alternative views on |
|
1115 |
surjectivity. These rules are already declared as set or type rules |
|
1116 |
for the generic @{method cases} and @{method induct} methods, |
|
1117 |
respectively. |
|
1118 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
1119 |
\end{description} |
42908 | 1120 |
*} |
1121 |
||
1122 |
subsubsection {* Examples *} |
|
1123 |
||
1124 |
text {* Type definitions permit the introduction of abstract data |
|
1125 |
types in a safe way, namely by providing models based on already |
|
1126 |
existing types. Given some abstract axiomatic description @{text P} |
|
1127 |
of a type, this involves two steps: |
|
1128 |
||
1129 |
\begin{enumerate} |
|
1130 |
||
1131 |
\item Find an appropriate type @{text \<tau>} and subset @{text A} which |
|
1132 |
has the desired properties @{text P}, and make a type definition |
|
1133 |
based on this representation. |
|
1134 |
||
1135 |
\item Prove that @{text P} holds for @{text \<tau>} by lifting @{text P} |
|
1136 |
from the representation. |
|
26849 | 1137 |
|
42908 | 1138 |
\end{enumerate} |
1139 |
||
1140 |
You can later forget about the representation and work solely in |
|
1141 |
terms of the abstract properties @{text P}. |
|
1142 |
||
1143 |
\medskip The following trivial example pulls a three-element type |
|
1144 |
into existence within the formal logical environment of HOL. *} |
|
1145 |
||
49834 | 1146 |
typedef three = "{(True, True), (True, False), (False, True)}" |
42908 | 1147 |
by blast |
1148 |
||
1149 |
definition "One = Abs_three (True, True)" |
|
1150 |
definition "Two = Abs_three (True, False)" |
|
1151 |
definition "Three = Abs_three (False, True)" |
|
1152 |
||
1153 |
lemma three_distinct: "One \<noteq> Two" "One \<noteq> Three" "Two \<noteq> Three" |
|
49812
e3945ddcb9aa
eliminated some remaining uses of typedef with implicit set definition;
wenzelm
parents:
48985
diff
changeset
|
1154 |
by (simp_all add: One_def Two_def Three_def Abs_three_inject) |
42908 | 1155 |
|
1156 |
lemma three_cases: |
|
1157 |
fixes x :: three obtains "x = One" | "x = Two" | "x = Three" |
|
49812
e3945ddcb9aa
eliminated some remaining uses of typedef with implicit set definition;
wenzelm
parents:
48985
diff
changeset
|
1158 |
by (cases x) (auto simp: One_def Two_def Three_def Abs_three_inject) |
42908 | 1159 |
|
1160 |
text {* Note that such trivial constructions are better done with |
|
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
1161 |
derived specification mechanisms such as @{command datatype_new}: *} |
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
1162 |
|
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
1163 |
datatype_new three' = One' | Two' | Three' |
42908 | 1164 |
|
1165 |
text {* This avoids re-doing basic definitions and proofs from the |
|
1166 |
primitive @{command typedef} above. *} |
|
26849 | 1167 |
|
1168 |
||
50109 | 1169 |
|
41396 | 1170 |
section {* Functorial structure of types *} |
1171 |
||
1172 |
text {* |
|
1173 |
\begin{matharray}{rcl} |
|
55467
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
blanchet
parents:
55372
diff
changeset
|
1174 |
@{command_def (HOL) "functor"} & : & @{text "local_theory \<rightarrow> proof(prove)"} |
41396 | 1175 |
\end{matharray} |
1176 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1177 |
@{rail \<open> |
55467
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
blanchet
parents:
55372
diff
changeset
|
1178 |
@@{command (HOL) functor} (@{syntax name} ':')? @{syntax term} |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1179 |
\<close>} |
41396 | 1180 |
|
1181 |
\begin{description} |
|
1182 |
||
55467
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
blanchet
parents:
55372
diff
changeset
|
1183 |
\item @{command (HOL) "functor"}~@{text "prefix: m"} allows to |
42617 | 1184 |
prove and register properties about the functorial structure of type |
1185 |
constructors. These properties then can be used by other packages |
|
1186 |
to deal with those type constructors in certain type constructions. |
|
1187 |
Characteristic theorems are noted in the current local theory. By |
|
1188 |
default, they are prefixed with the base name of the type |
|
1189 |
constructor, an explicit prefix can be given alternatively. |
|
41396 | 1190 |
|
1191 |
The given term @{text "m"} is considered as \emph{mapper} for the |
|
1192 |
corresponding type constructor and must conform to the following |
|
1193 |
type pattern: |
|
1194 |
||
1195 |
\begin{matharray}{lll} |
|
1196 |
@{text "m"} & @{text "::"} & |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1197 |
@{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>k \<Rightarrow> (\<^vec>\<alpha>\<^sub>n) t \<Rightarrow> (\<^vec>\<beta>\<^sub>n) t"} \\ |
41396 | 1198 |
\end{matharray} |
1199 |
||
1200 |
\noindent where @{text t} is the type constructor, @{text |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1201 |
"\<^vec>\<alpha>\<^sub>n"} and @{text "\<^vec>\<beta>\<^sub>n"} are distinct |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1202 |
type variables free in the local theory and @{text "\<sigma>\<^sub>1"}, |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1203 |
\ldots, @{text "\<sigma>\<^sub>k"} is a subsequence of @{text "\<alpha>\<^sub>1 \<Rightarrow> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1204 |
\<beta>\<^sub>1"}, @{text "\<beta>\<^sub>1 \<Rightarrow> \<alpha>\<^sub>1"}, \ldots, |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1205 |
@{text "\<alpha>\<^sub>n \<Rightarrow> \<beta>\<^sub>n"}, @{text "\<beta>\<^sub>n \<Rightarrow> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1206 |
\<alpha>\<^sub>n"}. |
41396 | 1207 |
|
1208 |
\end{description} |
|
1209 |
*} |
|
1210 |
||
47859 | 1211 |
|
50109 | 1212 |
section {* Quotient types *} |
1213 |
||
1214 |
text {* |
|
1215 |
\begin{matharray}{rcl} |
|
1216 |
@{command_def (HOL) "quotient_type"} & : & @{text "local_theory \<rightarrow> proof(prove)"}\\ |
|
1217 |
@{command_def (HOL) "quotient_definition"} & : & @{text "local_theory \<rightarrow> proof(prove)"}\\ |
|
1218 |
@{command_def (HOL) "print_quotmapsQ3"} & : & @{text "context \<rightarrow>"}\\ |
|
1219 |
@{command_def (HOL) "print_quotientsQ3"} & : & @{text "context \<rightarrow>"}\\ |
|
1220 |
@{command_def (HOL) "print_quotconsts"} & : & @{text "context \<rightarrow>"}\\ |
|
1221 |
@{method_def (HOL) "lifting"} & : & @{text method} \\ |
|
1222 |
@{method_def (HOL) "lifting_setup"} & : & @{text method} \\ |
|
1223 |
@{method_def (HOL) "descending"} & : & @{text method} \\ |
|
1224 |
@{method_def (HOL) "descending_setup"} & : & @{text method} \\ |
|
1225 |
@{method_def (HOL) "partiality_descending"} & : & @{text method} \\ |
|
1226 |
@{method_def (HOL) "partiality_descending_setup"} & : & @{text method} \\ |
|
1227 |
@{method_def (HOL) "regularize"} & : & @{text method} \\ |
|
1228 |
@{method_def (HOL) "injection"} & : & @{text method} \\ |
|
1229 |
@{method_def (HOL) "cleaning"} & : & @{text method} \\ |
|
1230 |
@{attribute_def (HOL) "quot_thm"} & : & @{text attribute} \\ |
|
1231 |
@{attribute_def (HOL) "quot_lifted"} & : & @{text attribute} \\ |
|
1232 |
@{attribute_def (HOL) "quot_respect"} & : & @{text attribute} \\ |
|
1233 |
@{attribute_def (HOL) "quot_preserve"} & : & @{text attribute} \\ |
|
1234 |
\end{matharray} |
|
1235 |
||
1236 |
The quotient package defines a new quotient type given a raw type |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1237 |
and a partial equivalence relation. The package also historically |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1238 |
includes automation for transporting definitions and theorems. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1239 |
But most of this automation was superseded by the Lifting and Transfer |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1240 |
packages. The user should consider using these two new packages for |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1241 |
lifting definitions and transporting theorems. |
50109 | 1242 |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1243 |
@{rail \<open> |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1244 |
@@{command (HOL) quotient_type} (spec) |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1245 |
; |
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
1246 |
spec: @{syntax typespec} @{syntax mixfix}? '=' \<newline> |
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
1247 |
@{syntax type} '/' ('partial' ':')? @{syntax term} \<newline> |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1248 |
(@'morphisms' @{syntax name} @{syntax name})? (@'parametric' @{syntax thmref})? |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1249 |
\<close>} |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1250 |
|
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1251 |
@{rail \<open> |
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
1252 |
@@{command (HOL) quotient_definition} constdecl? @{syntax thmdecl}? \<newline> |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1253 |
@{syntax term} 'is' @{syntax term} |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1254 |
; |
50109 | 1255 |
constdecl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1256 |
\<close>} |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1257 |
|
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1258 |
@{rail \<open> |
50109 | 1259 |
@@{method (HOL) lifting} @{syntax thmrefs}? |
1260 |
; |
|
1261 |
@@{method (HOL) lifting_setup} @{syntax thmrefs}? |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1262 |
\<close>} |
50109 | 1263 |
|
1264 |
\begin{description} |
|
1265 |
||
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1266 |
\item @{command (HOL) "quotient_type"} defines a new quotient type @{text \<tau>}. The |
50109 | 1267 |
injection from a quotient type to a raw type is called @{text |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1268 |
rep_\<tau>}, its inverse @{text abs_\<tau>} unless explicit @{keyword (HOL) |
50109 | 1269 |
"morphisms"} specification provides alternative names. @{command |
1270 |
(HOL) "quotient_type"} requires the user to prove that the relation |
|
1271 |
is an equivalence relation (predicate @{text equivp}), unless the |
|
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1272 |
user specifies explicitly @{text partial} in which case the |
50109 | 1273 |
obligation is @{text part_equivp}. A quotient defined with @{text |
1274 |
partial} is weaker in the sense that less things can be proved |
|
1275 |
automatically. |
|
1276 |
||
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1277 |
The command internally proves a Quotient theorem and sets up the Lifting |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1278 |
package by the command @{command (HOL) setup_lifting}. Thus the Lifting |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1279 |
and Transfer packages can be used also with quotient types defined by |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1280 |
@{command (HOL) "quotient_type"} without any extra set-up. The parametricity |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1281 |
theorem for the equivalence relation R can be provided as an extra argument |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1282 |
of the command and is passed to the corresponding internal call of @{command (HOL) setup_lifting}. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1283 |
This theorem allows the Lifting package to generate a stronger transfer rule for equality. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1284 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1285 |
\end{description} |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1286 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1287 |
The most of the rest of the package was superseded by the Lifting and Transfer |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1288 |
packages. The user should consider using these two new packages for |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1289 |
lifting definitions and transporting theorems. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1290 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1291 |
\begin{description} |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1292 |
|
50109 | 1293 |
\item @{command (HOL) "quotient_definition"} defines a constant on |
1294 |
the quotient type. |
|
1295 |
||
1296 |
\item @{command (HOL) "print_quotmapsQ3"} prints quotient map |
|
1297 |
functions. |
|
1298 |
||
1299 |
\item @{command (HOL) "print_quotientsQ3"} prints quotients. |
|
1300 |
||
1301 |
\item @{command (HOL) "print_quotconsts"} prints quotient constants. |
|
1302 |
||
1303 |
\item @{method (HOL) "lifting"} and @{method (HOL) "lifting_setup"} |
|
1304 |
methods match the current goal with the given raw theorem to be |
|
1305 |
lifted producing three new subgoals: regularization, injection and |
|
1306 |
cleaning subgoals. @{method (HOL) "lifting"} tries to apply the |
|
1307 |
heuristics for automatically solving these three subgoals and |
|
1308 |
leaves only the subgoals unsolved by the heuristics to the user as |
|
1309 |
opposed to @{method (HOL) "lifting_setup"} which leaves the three |
|
1310 |
subgoals unsolved. |
|
1311 |
||
1312 |
\item @{method (HOL) "descending"} and @{method (HOL) |
|
1313 |
"descending_setup"} try to guess a raw statement that would lift |
|
1314 |
to the current subgoal. Such statement is assumed as a new subgoal |
|
1315 |
and @{method (HOL) "descending"} continues in the same way as |
|
1316 |
@{method (HOL) "lifting"} does. @{method (HOL) "descending"} tries |
|
1317 |
to solve the arising regularization, injection and cleaning |
|
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1318 |
subgoals with the analogous method @{method (HOL) |
50109 | 1319 |
"descending_setup"} which leaves the four unsolved subgoals. |
1320 |
||
1321 |
\item @{method (HOL) "partiality_descending"} finds the regularized |
|
1322 |
theorem that would lift to the current subgoal, lifts it and |
|
1323 |
leaves as a subgoal. This method can be used with partial |
|
1324 |
equivalence quotients where the non regularized statements would |
|
1325 |
not be true. @{method (HOL) "partiality_descending_setup"} leaves |
|
1326 |
the injection and cleaning subgoals unchanged. |
|
1327 |
||
1328 |
\item @{method (HOL) "regularize"} applies the regularization |
|
1329 |
heuristics to the current subgoal. |
|
1330 |
||
1331 |
\item @{method (HOL) "injection"} applies the injection heuristics |
|
1332 |
to the current goal using the stored quotient respectfulness |
|
1333 |
theorems. |
|
1334 |
||
1335 |
\item @{method (HOL) "cleaning"} applies the injection cleaning |
|
1336 |
heuristics to the current subgoal using the stored quotient |
|
1337 |
preservation theorems. |
|
1338 |
||
1339 |
\item @{attribute (HOL) quot_lifted} attribute tries to |
|
1340 |
automatically transport the theorem to the quotient type. |
|
1341 |
The attribute uses all the defined quotients types and quotient |
|
1342 |
constants often producing undesired results or theorems that |
|
1343 |
cannot be lifted. |
|
1344 |
||
1345 |
\item @{attribute (HOL) quot_respect} and @{attribute (HOL) |
|
1346 |
quot_preserve} attributes declare a theorem as a respectfulness |
|
1347 |
and preservation theorem respectively. These are stored in the |
|
1348 |
local theory store and used by the @{method (HOL) "injection"} |
|
1349 |
and @{method (HOL) "cleaning"} methods respectively. |
|
1350 |
||
1351 |
\item @{attribute (HOL) quot_thm} declares that a certain theorem |
|
1352 |
is a quotient extension theorem. Quotient extension theorems |
|
1353 |
allow for quotienting inside container types. Given a polymorphic |
|
1354 |
type that serves as a container, a map function defined for this |
|
55467
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
blanchet
parents:
55372
diff
changeset
|
1355 |
container using @{command (HOL) "functor"} and a relation |
50109 | 1356 |
map defined for for the container type, the quotient extension |
1357 |
theorem should be @{term "Quotient3 R Abs Rep \<Longrightarrow> Quotient3 |
|
1358 |
(rel_map R) (map Abs) (map Rep)"}. Quotient extension theorems |
|
1359 |
are stored in a database and are used all the steps of lifting |
|
1360 |
theorems. |
|
1361 |
||
1362 |
\end{description} |
|
1363 |
*} |
|
1364 |
||
1365 |
||
1366 |
section {* Definition by specification \label{sec:hol-specification} *} |
|
1367 |
||
1368 |
text {* |
|
1369 |
\begin{matharray}{rcl} |
|
1370 |
@{command_def (HOL) "specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\ |
|
1371 |
\end{matharray} |
|
1372 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1373 |
@{rail \<open> |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
55944
diff
changeset
|
1374 |
@@{command (HOL) specification} '(' (decl +) ')' \<newline> |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
55944
diff
changeset
|
1375 |
(@{syntax thmdecl}? @{syntax prop} +) |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1376 |
; |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
55944
diff
changeset
|
1377 |
decl: (@{syntax name} ':')? @{syntax term} ('(' @'overloaded' ')')? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1378 |
\<close>} |
50109 | 1379 |
|
1380 |
\begin{description} |
|
1381 |
||
1382 |
\item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a |
|
1383 |
goal stating the existence of terms with the properties specified to |
|
1384 |
hold for the constants given in @{text decls}. After finishing the |
|
1385 |
proof, the theory will be augmented with definitions for the given |
|
1386 |
constants, as well as with theorems stating the properties for these |
|
1387 |
constants. |
|
1388 |
||
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
55944
diff
changeset
|
1389 |
@{text decl} declares a constant to be defined by the |
50109 | 1390 |
specification given. The definition for the constant @{text c} is |
1391 |
bound to the name @{text c_def} unless a theorem name is given in |
|
1392 |
the declaration. Overloaded constants should be declared as such. |
|
1393 |
||
1394 |
\end{description} |
|
1395 |
*} |
|
1396 |
||
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
55944
diff
changeset
|
1397 |
|
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1398 |
section {* Adhoc overloading of constants *} |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1399 |
|
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1400 |
text {* |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1401 |
\begin{tabular}{rcll} |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1402 |
@{command_def "adhoc_overloading"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1403 |
@{command_def "no_adhoc_overloading"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1404 |
@{attribute_def "show_variants"} & : & @{text "attribute"} & default @{text false} \\ |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1405 |
\end{tabular} |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1406 |
|
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1407 |
\medskip |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1408 |
|
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1409 |
Adhoc overloading allows to overload a constant depending on |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1410 |
its type. Typically this involves the introduction of an |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1411 |
uninterpreted constant (used for input and output) and the addition |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1412 |
of some variants (used internally). For examples see |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1413 |
@{file "~~/src/HOL/ex/Adhoc_Overloading_Examples.thy"} and |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1414 |
@{file "~~/src/HOL/Library/Monad_Syntax.thy"}. |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1415 |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1416 |
@{rail \<open> |
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1417 |
(@@{command adhoc_overloading} | @@{command no_adhoc_overloading}) |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1418 |
(@{syntax nameref} (@{syntax term} + ) + @'and') |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1419 |
\<close>} |
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1420 |
|
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1421 |
\begin{description} |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1422 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1423 |
\item @{command "adhoc_overloading"}~@{text "c v\<^sub>1 ... v\<^sub>n"} |
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1424 |
associates variants with an existing constant. |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1425 |
|
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1426 |
\item @{command "no_adhoc_overloading"} is similar to |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1427 |
@{command "adhoc_overloading"}, but removes the specified variants |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1428 |
from the present context. |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1429 |
|
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1430 |
\item @{attribute "show_variants"} controls printing of variants |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1431 |
of overloaded constants. If enabled, the internally used variants |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1432 |
are printed instead of their respective overloaded constants. This |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1433 |
is occasionally useful to check whether the system agrees with a |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1434 |
user's expectations about derived variants. |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1435 |
|
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1436 |
\end{description} |
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
1437 |
*} |
50109 | 1438 |
|
57507 | 1439 |
|
50109 | 1440 |
chapter {* Proof tools *} |
1441 |
||
1442 |
section {* Adhoc tuples *} |
|
1443 |
||
1444 |
text {* |
|
1445 |
\begin{matharray}{rcl} |
|
1446 |
@{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\ |
|
1447 |
\end{matharray} |
|
1448 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1449 |
@{rail \<open> |
50109 | 1450 |
@@{attribute (HOL) split_format} ('(' 'complete' ')')? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1451 |
\<close>} |
50109 | 1452 |
|
1453 |
\begin{description} |
|
1454 |
||
1455 |
\item @{attribute (HOL) split_format}\ @{text "(complete)"} causes |
|
1456 |
arguments in function applications to be represented canonically |
|
1457 |
according to their tuple type structure. |
|
1458 |
||
1459 |
Note that this operation tends to invent funny names for new local |
|
1460 |
parameters introduced. |
|
1461 |
||
1462 |
\end{description} |
|
1463 |
*} |
|
1464 |
||
1465 |
||
47821 | 1466 |
section {* Transfer package *} |
1467 |
||
1468 |
text {* |
|
1469 |
\begin{matharray}{rcl} |
|
1470 |
@{method_def (HOL) "transfer"} & : & @{text method} \\ |
|
1471 |
@{method_def (HOL) "transfer'"} & : & @{text method} \\ |
|
1472 |
@{method_def (HOL) "transfer_prover"} & : & @{text method} \\ |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1473 |
@{attribute_def (HOL) "Transfer.transferred"} & : & @{text attribute} \\ |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1474 |
@{attribute_def (HOL) "untransferred"} & : & @{text attribute} \\ |
47821 | 1475 |
@{attribute_def (HOL) "transfer_rule"} & : & @{text attribute} \\ |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1476 |
@{attribute_def (HOL) "transfer_domain_rule"} & : & @{text attribute} \\ |
47821 | 1477 |
@{attribute_def (HOL) "relator_eq"} & : & @{text attribute} \\ |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1478 |
@{attribute_def (HOL) "relator_domain"} & : & @{text attribute} \\ |
47821 | 1479 |
\end{matharray} |
1480 |
||
1481 |
\begin{description} |
|
1482 |
||
1483 |
\item @{method (HOL) "transfer"} method replaces the current subgoal |
|
1484 |
with a logically equivalent one that uses different types and |
|
1485 |
constants. The replacement of types and constants is guided by the |
|
1486 |
database of transfer rules. Goals are generalized over all free |
|
1487 |
variables by default; this is necessary for variables whose types |
|
1488 |
change, but can be overridden for specific variables with e.g. |
|
1489 |
@{text "transfer fixing: x y z"}. |
|
1490 |
||
1491 |
\item @{method (HOL) "transfer'"} is a variant of @{method (HOL) |
|
1492 |
transfer} that allows replacing a subgoal with one that is |
|
1493 |
logically stronger (rather than equivalent). For example, a |
|
1494 |
subgoal involving equality on a quotient type could be replaced |
|
1495 |
with a subgoal involving equality (instead of the corresponding |
|
1496 |
equivalence relation) on the underlying raw type. |
|
1497 |
||
1498 |
\item @{method (HOL) "transfer_prover"} method assists with proving |
|
1499 |
a transfer rule for a new constant, provided the constant is |
|
1500 |
defined in terms of other constants that already have transfer |
|
1501 |
rules. It should be applied after unfolding the constant |
|
1502 |
definitions. |
|
1503 |
||
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1504 |
\item @{attribute (HOL) "untransferred"} proves the same equivalent theorem |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1505 |
as @{method (HOL) "transfer"} internally does. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1506 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1507 |
\item @{attribute (HOL) Transfer.transferred} works in the opposite |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1508 |
direction than @{method (HOL) "transfer'"}. E.g., given the transfer |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1509 |
relation @{text "ZN x n \<equiv> (x = int n)"}, corresponding transfer rules and the theorem |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1510 |
@{text "\<forall>x::int \<in> {0..}. x < x + 1"}, the attribute would prove |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1511 |
@{text "\<forall>n::nat. n < n + 1"}. The attribute is still in experimental |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1512 |
phase of development. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1513 |
|
47821 | 1514 |
\item @{attribute (HOL) "transfer_rule"} attribute maintains a |
1515 |
collection of transfer rules, which relate constants at two |
|
1516 |
different types. Typical transfer rules may relate different type |
|
1517 |
instances of the same polymorphic constant, or they may relate an |
|
1518 |
operation on a raw type to a corresponding operation on an |
|
1519 |
abstract type (quotient or subtype). For example: |
|
1520 |
||
1521 |
@{text "((A ===> B) ===> list_all2 A ===> list_all2 B) map map"}\\ |
|
1522 |
@{text "(cr_int ===> cr_int ===> cr_int) (\<lambda>(x,y) (u,v). (x+u, y+v)) plus"} |
|
1523 |
||
1524 |
Lemmas involving predicates on relations can also be registered |
|
1525 |
using the same attribute. For example: |
|
1526 |
||
1527 |
@{text "bi_unique A \<Longrightarrow> (list_all2 A ===> op =) distinct distinct"}\\ |
|
55944 | 1528 |
@{text "\<lbrakk>bi_unique A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (rel_prod A B)"} |
47821 | 1529 |
|
57829 | 1530 |
Preservation of predicates on relations (@{text "bi_unique, bi_total, |
1531 |
right_unique, right_total, left_unique, left_total"}) with the respect to a relator |
|
1532 |
is proved automatically if the involved type is BNF\cite{isabelle-datatypes} without dead variables. |
|
1533 |
||
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1534 |
\item @{attribute (HOL) "transfer_domain_rule"} attribute maintains a collection |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1535 |
of rules, which specify a domain of a transfer relation by a predicate. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1536 |
E.g., given the transfer relation @{text "ZN x n \<equiv> (x = int n)"}, |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1537 |
one can register the following transfer domain rule: |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1538 |
@{text "Domainp ZN = (\<lambda>x. x \<ge> 0)"}. The rules allow the package to produce |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1539 |
more readable transferred goals, e.g., when quantifiers are transferred. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1540 |
|
47821 | 1541 |
\item @{attribute (HOL) relator_eq} attribute collects identity laws |
57829 | 1542 |
for relators of various type constructors, e.g. @{term "rel_set |
47821 | 1543 |
(op =) = (op =)"}. The @{method (HOL) transfer} method uses these |
1544 |
lemmas to infer transfer rules for non-polymorphic constants on |
|
57829 | 1545 |
the fly. For examples see @{file |
1546 |
"~~/src/HOL/Lifting_Set.thy"} or @{file "~~/src/HOL/Lifting.thy"}. |
|
1547 |
This property is proved automatically if the involved type is BNF without dead variables. |
|
47821 | 1548 |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1549 |
\item @{attribute_def (HOL) "relator_domain"} attribute collects rules |
57829 | 1550 |
describing domains of relators by predicators. E.g., |
1551 |
@{term "Domainp (rel_set T) = (\<lambda>A. Ball A (Domainp T))"}. This allows the package |
|
1552 |
to lift transfer domain rules through type constructors. For examples see @{file |
|
1553 |
"~~/src/HOL/Lifting_Set.thy"} or @{file "~~/src/HOL/Lifting.thy"}. |
|
1554 |
This property is proved automatically if the involved type is BNF without dead variables. |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1555 |
|
47821 | 1556 |
\end{description} |
1557 |
||
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1558 |
Theoretical background can be found in \cite{Huffman-Kuncar:2013:lifting_transfer}. |
47821 | 1559 |
*} |
1560 |
||
47859 | 1561 |
|
47802 | 1562 |
section {* Lifting package *} |
1563 |
||
1564 |
text {* |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1565 |
The Lifting package allows users to lift terms of the raw type to the abstract type, which is |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1566 |
a necessary step in building a library for an abstract type. Lifting defines a new constant |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1567 |
by combining coercion functions (Abs and Rep) with the raw term. It also proves an appropriate |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1568 |
transfer rule for the Transfer package and, if possible, an equation for the code generator. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1569 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1570 |
The Lifting package provides two main commands: @{command (HOL) "setup_lifting"} for initializing |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1571 |
the package to work with a new type, and @{command (HOL) "lift_definition"} for lifting constants. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1572 |
The Lifting package works with all four kinds of type abstraction: type copies, subtypes, |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1573 |
total quotients and partial quotients. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1574 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1575 |
Theoretical background can be found in \cite{Huffman-Kuncar:2013:lifting_transfer}. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1576 |
|
47802 | 1577 |
\begin{matharray}{rcl} |
1578 |
@{command_def (HOL) "setup_lifting"} & : & @{text "local_theory \<rightarrow> local_theory"}\\ |
|
1579 |
@{command_def (HOL) "lift_definition"} & : & @{text "local_theory \<rightarrow> proof(prove)"}\\ |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1580 |
@{command_def (HOL) "lifting_forget"} & : & @{text "local_theory \<rightarrow> local_theory"}\\ |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1581 |
@{command_def (HOL) "lifting_update"} & : & @{text "local_theory \<rightarrow> local_theory"}\\ |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
53015
diff
changeset
|
1582 |
@{command_def (HOL) "print_quot_maps"} & : & @{text "context \<rightarrow>"}\\ |
47802 | 1583 |
@{command_def (HOL) "print_quotients"} & : & @{text "context \<rightarrow>"}\\ |
1584 |
@{attribute_def (HOL) "quot_map"} & : & @{text attribute} \\ |
|
56519
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents:
56518
diff
changeset
|
1585 |
@{attribute_def (HOL) "relator_eq_onp"} & : & @{text attribute} \\ |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1586 |
@{attribute_def (HOL) "relator_mono"} & : & @{text attribute} \\ |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1587 |
@{attribute_def (HOL) "relator_distr"} & : & @{text attribute} \\ |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1588 |
@{attribute_def (HOL) "quot_del"} & : & @{text attribute} \\ |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1589 |
@{attribute_def (HOL) "lifting_restore"} & : & @{text attribute} \\ |
47802 | 1590 |
\end{matharray} |
1591 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1592 |
@{rail \<open> |
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
1593 |
@@{command (HOL) setup_lifting} ('(' 'no_code' ')')? \<newline> |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1594 |
@{syntax thmref} @{syntax thmref}? (@'parametric' @{syntax thmref})?; |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1595 |
\<close>} |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1596 |
|
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1597 |
@{rail \<open> |
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
1598 |
@@{command (HOL) lift_definition} @{syntax name} '::' @{syntax type} @{syntax mixfix}? \<newline> |
57829 | 1599 |
'is' @{syntax term} (@'parametric' (@{syntax thmref}+))?; |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1600 |
\<close>} |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1601 |
|
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1602 |
@{rail \<open> |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1603 |
@@{command (HOL) lifting_forget} @{syntax nameref}; |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1604 |
\<close>} |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1605 |
|
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1606 |
@{rail \<open> |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1607 |
@@{command (HOL) lifting_update} @{syntax nameref}; |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1608 |
\<close>} |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1609 |
|
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1610 |
@{rail \<open> |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1611 |
@@{attribute (HOL) lifting_restore} @{syntax thmref} (@{syntax thmref} @{syntax thmref})?; |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1612 |
\<close>} |
47802 | 1613 |
|
1614 |
\begin{description} |
|
1615 |
||
47859 | 1616 |
\item @{command (HOL) "setup_lifting"} Sets up the Lifting package |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1617 |
to work with a user-defined type. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1618 |
The command supports two modes. The first one is a low-level mode when |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1619 |
the user must provide as a first |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1620 |
argument of @{command (HOL) "setup_lifting"} a |
57829 | 1621 |
quotient theorem @{term "Quotient R Abs Rep T"}. The |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1622 |
package configures a transfer rule for equality, a domain transfer |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1623 |
rules and sets up the @{command_def (HOL) "lift_definition"} |
57829 | 1624 |
command to work with the abstract type. An optional theorem @{term "reflp R"}, which certifies that |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1625 |
the equivalence relation R is total, |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1626 |
can be provided as a second argument. This allows the package to generate stronger transfer |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1627 |
rules. And finally, the parametricity theorem for R can be provided as a third argument. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1628 |
This allows the package to generate a stronger transfer rule for equality. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1629 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1630 |
Users generally will not prove the @{text Quotient} theorem manually for |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1631 |
new types, as special commands exist to automate the process. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1632 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1633 |
When a new subtype is defined by @{command (HOL) typedef}, @{command (HOL) "lift_definition"} |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1634 |
can be used in its |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1635 |
second mode, where only the type_definition theorem @{text "type_definition Rep Abs A"} |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1636 |
is used as an argument of the command. The command internally proves the corresponding |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1637 |
Quotient theorem and registers it with @{command (HOL) setup_lifting} using its first mode. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1638 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1639 |
For quotients, the command @{command (HOL) quotient_type} can be used. The command defines |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1640 |
a new quotient type and similarly to the previous case, the corresponding Quotient theorem is proved |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1641 |
and registered by @{command (HOL) setup_lifting}. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1642 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1643 |
The command @{command (HOL) "setup_lifting"} also sets up the code generator |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1644 |
for the new type. Later on, when a new constant is defined by @{command (HOL) "lift_definition"}, |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1645 |
the Lifting package proves and registers a code equation (if there is one) for the new constant. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1646 |
If the option @{text "no_code"} is specified, the Lifting package does not set up the code |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1647 |
generator and as a consequence no code equations involving an abstract type are registered |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1648 |
by @{command (HOL) "lift_definition"}. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1649 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1650 |
\item @{command (HOL) "lift_definition"} @{text "f :: \<tau>"} @{keyword (HOL) "is"} @{text t} |
47859 | 1651 |
Defines a new function @{text f} with an abstract type @{text \<tau>} |
1652 |
in terms of a corresponding operation @{text t} on a |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1653 |
representation type. More formally, if @{text "t :: \<sigma>"}, then |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1654 |
the command builds a term @{text "F"} as a corresponding combination of abstraction |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1655 |
and representation functions such that @{text "F :: \<sigma> \<Rightarrow> \<tau>" } and |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1656 |
defines @{text f} is as @{text "f \<equiv> F t"}. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1657 |
The term @{text t} does not have to be necessarily a constant but it can be any term. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1658 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1659 |
The command opens a proof environment and the user must discharge |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1660 |
a respectfulness proof obligation. For a type copy, i.e., a typedef with @{text |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1661 |
UNIV}, the obligation is discharged automatically. The proof goal is |
47802 | 1662 |
presented in a user-friendly, readable form. A respectfulness |
47859 | 1663 |
theorem in the standard format @{text f.rsp} and a transfer rule |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1664 |
@{text f.transfer} for the Transfer package are generated by the |
47859 | 1665 |
package. |
47802 | 1666 |
|
57829 | 1667 |
The user can specify a parametricity theorems for @{text t} after the keyword |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1668 |
@{keyword "parametric"}, which allows the command |
57829 | 1669 |
to generate parametric transfer rules for @{text f}. |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1670 |
|
50879 | 1671 |
For each constant defined through trivial quotients (type copies or |
1672 |
subtypes) @{text f.rep_eq} is generated. The equation is a code certificate |
|
1673 |
that defines @{text f} using the representation function. |
|
1674 |
||
1675 |
For each constant @{text f.abs_eq} is generated. The equation is unconditional |
|
1676 |
for total quotients. The equation defines @{text f} using |
|
1677 |
the abstraction function. |
|
1678 |
||
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1679 |
Integration with [@{attribute code} abstract]: For subtypes (e.g., |
47859 | 1680 |
corresponding to a datatype invariant, such as dlist), @{command |
50879 | 1681 |
(HOL) "lift_definition"} uses a code certificate theorem |
1682 |
@{text f.rep_eq} as a code equation. |
|
1683 |
||
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1684 |
Integration with [@{attribute code} equation]: For total quotients, @{command |
50879 | 1685 |
(HOL) "lift_definition"} uses @{text f.abs_eq} as a code equation. |
47802 | 1686 |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1687 |
\item @{command (HOL) lifting_forget} and @{command (HOL) lifting_update} |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1688 |
These two commands serve for storing and deleting the set-up of |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1689 |
the Lifting package and corresponding transfer rules defined by this package. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1690 |
This is useful for hiding of type construction details of an abstract type |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1691 |
when the construction is finished but it still allows additions to this construction |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1692 |
when this is later necessary. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1693 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1694 |
Whenever the Lifting package is set up with a new abstract type @{text "\<tau>"} by |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1695 |
@{command_def (HOL) "lift_definition"}, the package defines a new bundle |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1696 |
that is called @{text "\<tau>.lifting"}. This bundle already includes set-up for the Lifting package. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1697 |
The new transfer rules |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1698 |
introduced by @{command (HOL) "lift_definition"} can be stored in the bundle by |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1699 |
the command @{command (HOL) "lifting_update"} @{text "\<tau>.lifting"}. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1700 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1701 |
The command @{command (HOL) "lifting_forget"} @{text "\<tau>.lifting"} deletes set-up of the Lifting |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1702 |
package |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1703 |
for @{text \<tau>} and deletes all the transfer rules that were introduced |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1704 |
by @{command (HOL) "lift_definition"} using @{text \<tau>} as an abstract type. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1705 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1706 |
The stored set-up in a bundle can be reintroduced by the Isar commands for including a bundle |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1707 |
(@{command "include"}, @{keyword "includes"} and @{command "including"}). |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1708 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
53015
diff
changeset
|
1709 |
\item @{command (HOL) "print_quot_maps"} prints stored quotient map |
47859 | 1710 |
theorems. |
1711 |
||
1712 |
\item @{command (HOL) "print_quotients"} prints stored quotient |
|
1713 |
theorems. |
|
1714 |
||
1715 |
\item @{attribute (HOL) quot_map} registers a quotient map |
|
57829 | 1716 |
theorem, a theorem showing how to "lift" quotients over type constructors. |
1717 |
E.g., @{term "Quotient R Abs Rep T \<Longrightarrow> |
|
1718 |
Quotient (rel_set R) (image Abs) (image Rep) (rel_set T)"}. |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1719 |
For examples see @{file |
57829 | 1720 |
"~~/src/HOL/Lifting_Set.thy"} or @{file "~~/src/HOL/Lifting.thy"}. |
1721 |
This property is proved automatically if the involved type is BNF without dead variables. |
|
47859 | 1722 |
|
56519
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents:
56518
diff
changeset
|
1723 |
\item @{attribute (HOL) relator_eq_onp} registers a theorem that |
57829 | 1724 |
shows that a relator applied to an equality restricted by a predicate @{term P} (i.e., @{term |
1725 |
"eq_onp P"}) is equal |
|
1726 |
to a predicator applied to the @{term P}. The combinator @{const eq_onp} is used for |
|
1727 |
internal encoding of proper subtypes. Such theorems allows the package to hide @{text |
|
56519
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents:
56518
diff
changeset
|
1728 |
eq_onp} from a user in a user-readable form of a |
47859 | 1729 |
respectfulness theorem. For examples see @{file |
57829 | 1730 |
"~~/src/HOL/Lifting_Set.thy"} or @{file "~~/src/HOL/Lifting.thy"}. |
1731 |
This property is proved automatically if the involved type is BNF without dead variables. |
|
47802 | 1732 |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1733 |
\item @{attribute (HOL) "relator_mono"} registers a property describing a monotonicity of a relator. |
57829 | 1734 |
E.g., @{term "A \<le> B \<Longrightarrow> rel_set A \<le> rel_set B"}. |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1735 |
This property is needed for proving a stronger transfer rule in @{command_def (HOL) "lift_definition"} |
57829 | 1736 |
when a parametricity theorem for the raw term is specified and also for the reflexivity prover. |
1737 |
For examples see @{file |
|
1738 |
"~~/src/HOL/Lifting_Set.thy"} or @{file "~~/src/HOL/Lifting.thy"}. |
|
1739 |
This property is proved automatically if the involved type is BNF without dead variables. |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1740 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1741 |
\item @{attribute (HOL) "relator_distr"} registers a property describing a distributivity |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1742 |
of the relation composition and a relator. E.g., |
57829 | 1743 |
@{text "rel_set R \<circ>\<circ> rel_set S = rel_set (R \<circ>\<circ> S)"}. |
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1744 |
This property is needed for proving a stronger transfer rule in @{command_def (HOL) "lift_definition"} |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1745 |
when a parametricity theorem for the raw term is specified. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1746 |
When this equality does not hold unconditionally (e.g., for the function type), the user can specified |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1747 |
each direction separately and also register multiple theorems with different set of assumptions. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1748 |
This attribute can be used only after the monotonicity property was already registered by |
57829 | 1749 |
@{attribute (HOL) "relator_mono"}. For examples see @{file |
1750 |
"~~/src/HOL/Lifting_Set.thy"} or @{file "~~/src/HOL/Lifting.thy"}. |
|
1751 |
This property is proved automatically if the involved type is BNF without dead variables. |
|
50877 | 1752 |
|
1753 |
\item @{attribute (HOL) quot_del} deletes a corresponding Quotient theorem |
|
1754 |
from the Lifting infrastructure and thus de-register the corresponding quotient. |
|
1755 |
This effectively causes that @{command (HOL) lift_definition} will not |
|
54334
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1756 |
do any lifting for the corresponding type. This attribute is rather used for low-level |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1757 |
manipulation with set-up of the Lifting package because @{command (HOL) lifting_forget} is |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1758 |
preferred for normal usage. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1759 |
|
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1760 |
\item @{attribute (HOL) lifting_restore} @{text "Quotient_thm pcr_def pcr_cr_eq_thm"} |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1761 |
registers the Quotient theorem @{text Quotient_thm} in the Lifting infrastructure |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1762 |
and thus sets up lifting for an abstract type @{text \<tau>} (that is defined by @{text Quotient_thm}). |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1763 |
Optional theorems @{text pcr_def} and @{text pcr_cr_eq_thm} can be specified to register |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1764 |
the parametrized |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1765 |
correspondence relation for @{text \<tau>}. E.g., for @{text "'a dlist"}, @{text pcr_def} is |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1766 |
@{text "pcr_dlist A \<equiv> list_all2 A \<circ>\<circ> cr_dlist"} and @{text pcr_cr_eq_thm} is |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1767 |
@{text "pcr_dlist op= = op="}. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1768 |
This attribute is rather used for low-level |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1769 |
manipulation with set-up of the Lifting package because using of the bundle @{text \<tau>.lifting} |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1770 |
together with the commands @{command (HOL) lifting_forget} and @{command (HOL) lifting_update} is |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1771 |
preferred for normal usage. |
409d7f7247f4
update documentation of Lifting/Transfer and Quotient
kuncar
parents:
54017
diff
changeset
|
1772 |
|
57829 | 1773 |
\item Integration with the BNF package\cite{isabelle-datatypes}: |
1774 |
As already mentioned, the theorems that are registered |
|
1775 |
by the following attributes are proved and registered automatically if the involved type |
|
1776 |
is BNF without dead variables: @{attribute (HOL) quot_map}, @{attribute (HOL) relator_eq_onp}, |
|
1777 |
@{attribute (HOL) "relator_mono"}, @{attribute (HOL) "relator_distr"}. Also the definition of a |
|
1778 |
relator and predicator is provided automatically. Moreover, if the BNF represents a datatype, |
|
1779 |
simplification rules for a predicator are again proved automatically. |
|
1780 |
||
47802 | 1781 |
\end{description} |
1782 |
*} |
|
1783 |
||
47859 | 1784 |
|
43994 | 1785 |
section {* Coercive subtyping *} |
1786 |
||
1787 |
text {* |
|
1788 |
\begin{matharray}{rcl} |
|
1789 |
@{attribute_def (HOL) coercion} & : & @{text attribute} \\ |
|
1790 |
@{attribute_def (HOL) coercion_enabled} & : & @{text attribute} \\ |
|
1791 |
@{attribute_def (HOL) coercion_map} & : & @{text attribute} \\ |
|
1792 |
\end{matharray} |
|
1793 |
||
46280 | 1794 |
Coercive subtyping allows the user to omit explicit type |
1795 |
conversions, also called \emph{coercions}. Type inference will add |
|
1796 |
them as necessary when parsing a term. See |
|
1797 |
\cite{traytel-berghofer-nipkow-2011} for details. |
|
1798 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1799 |
@{rail \<open> |
43994 | 1800 |
@@{attribute (HOL) coercion} (@{syntax term})? |
1801 |
; |
|
1802 |
@@{attribute (HOL) coercion_map} (@{syntax term})? |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1803 |
\<close>} |
43994 | 1804 |
|
1805 |
\begin{description} |
|
1806 |
||
1807 |
\item @{attribute (HOL) "coercion"}~@{text "f"} registers a new |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1808 |
coercion function @{text "f :: \<sigma>\<^sub>1 \<Rightarrow> \<sigma>\<^sub>2"} where @{text "\<sigma>\<^sub>1"} and |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1809 |
@{text "\<sigma>\<^sub>2"} are type constructors without arguments. Coercions are |
46280 | 1810 |
composed by the inference algorithm if needed. Note that the type |
1811 |
inference algorithm is complete only if the registered coercions |
|
1812 |
form a lattice. |
|
43994 | 1813 |
|
46280 | 1814 |
\item @{attribute (HOL) "coercion_map"}~@{text "map"} registers a |
1815 |
new map function to lift coercions through type constructors. The |
|
1816 |
function @{text "map"} must conform to the following type pattern |
|
43994 | 1817 |
|
1818 |
\begin{matharray}{lll} |
|
1819 |
@{text "map"} & @{text "::"} & |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1820 |
@{text "f\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> f\<^sub>n \<Rightarrow> (\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t \<Rightarrow> (\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n) t"} \\ |
43994 | 1821 |
\end{matharray} |
1822 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1823 |
where @{text "t"} is a type constructor and @{text "f\<^sub>i"} is of type |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
1824 |
@{text "\<alpha>\<^sub>i \<Rightarrow> \<beta>\<^sub>i"} or @{text "\<beta>\<^sub>i \<Rightarrow> \<alpha>\<^sub>i"}. Registering a map function |
46280 | 1825 |
overwrites any existing map function for this particular type |
1826 |
constructor. |
|
43994 | 1827 |
|
1828 |
\item @{attribute (HOL) "coercion_enabled"} enables the coercion |
|
1829 |
inference algorithm. |
|
1830 |
||
1831 |
\end{description} |
|
1832 |
*} |
|
1833 |
||
1834 |
||
26849 | 1835 |
section {* Arithmetic proof support *} |
1836 |
||
1837 |
text {* |
|
1838 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1839 |
@{method_def (HOL) arith} & : & @{text method} \\ |
30863 | 1840 |
@{attribute_def (HOL) arith} & : & @{text attribute} \\ |
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1841 |
@{attribute_def (HOL) arith_split} & : & @{text attribute} \\ |
26849 | 1842 |
\end{matharray} |
1843 |
||
46280 | 1844 |
\begin{description} |
26849 | 1845 |
|
46280 | 1846 |
\item @{method (HOL) arith} decides linear arithmetic problems (on |
1847 |
types @{text nat}, @{text int}, @{text real}). Any current facts |
|
1848 |
are inserted into the goal before running the procedure. |
|
26849 | 1849 |
|
46280 | 1850 |
\item @{attribute (HOL) arith} declares facts that are supplied to |
1851 |
the arithmetic provers implicitly. |
|
1852 |
||
1853 |
\item @{attribute (HOL) arith_split} attribute declares case split |
|
30865 | 1854 |
rules to be expanded before @{method (HOL) arith} is invoked. |
30863 | 1855 |
|
46280 | 1856 |
\end{description} |
1857 |
||
1858 |
Note that a simpler (but faster) arithmetic prover is already |
|
1859 |
invoked by the Simplifier. |
|
26849 | 1860 |
*} |
1861 |
||
1862 |
||
30169 | 1863 |
section {* Intuitionistic proof search *} |
1864 |
||
1865 |
text {* |
|
1866 |
\begin{matharray}{rcl} |
|
30171 | 1867 |
@{method_def (HOL) iprover} & : & @{text method} \\ |
30169 | 1868 |
\end{matharray} |
1869 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1870 |
@{rail \<open> |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1871 |
@@{method (HOL) iprover} (@{syntax rulemod} *) |
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1872 |
\<close>} |
30169 | 1873 |
|
46280 | 1874 |
\begin{description} |
1875 |
||
1876 |
\item @{method (HOL) iprover} performs intuitionistic proof search, |
|
1877 |
depending on specifically declared rules from the context, or given |
|
1878 |
as explicit arguments. Chained facts are inserted into the goal |
|
1879 |
before commencing proof search. |
|
35613 | 1880 |
|
30169 | 1881 |
Rules need to be classified as @{attribute (Pure) intro}, |
1882 |
@{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the |
|
1883 |
``@{text "!"}'' indicator refers to ``safe'' rules, which may be |
|
1884 |
applied aggressively (without considering back-tracking later). |
|
1885 |
Rules declared with ``@{text "?"}'' are ignored in proof search (the |
|
42626 | 1886 |
single-step @{method (Pure) rule} method still observes these). An |
30169 | 1887 |
explicit weight annotation may be given as well; otherwise the |
1888 |
number of rule premises will be taken into account here. |
|
46280 | 1889 |
|
1890 |
\end{description} |
|
30169 | 1891 |
*} |
1892 |
||
46280 | 1893 |
|
43578
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1894 |
section {* Model Elimination and Resolution *} |
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1895 |
|
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1896 |
text {* |
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1897 |
\begin{matharray}{rcl} |
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1898 |
@{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
|
1899 |
@{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
|
1900 |
\end{matharray} |
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1901 |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1902 |
@{rail \<open> |
43578
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1903 |
@@{method (HOL) meson} @{syntax thmrefs}? |
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1904 |
; |
46280 | 1905 |
@@{method (HOL) metis} |
1906 |
('(' ('partial_types' | 'full_types' | 'no_types' | @{syntax name}) ')')? |
|
1907 |
@{syntax thmrefs}? |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1908 |
\<close>} |
43578
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1909 |
|
46280 | 1910 |
\begin{description} |
1911 |
||
1912 |
\item @{method (HOL) meson} implements Loveland's model elimination |
|
1913 |
procedure \cite{loveland-78}. See @{file |
|
1914 |
"~~/src/HOL/ex/Meson_Test.thy"} for examples. |
|
43578
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1915 |
|
46280 | 1916 |
\item @{method (HOL) metis} combines ordered resolution and ordered |
1917 |
paramodulation to find first-order (or mildly higher-order) proofs. |
|
1918 |
The first optional argument specifies a type encoding; see the |
|
1919 |
Sledgehammer manual \cite{isabelle-sledgehammer} for details. The |
|
1920 |
directory @{file "~~/src/HOL/Metis_Examples"} contains several small |
|
1921 |
theories developed to a large extent using @{method (HOL) metis}. |
|
1922 |
||
1923 |
\end{description} |
|
43578
36ba44fe0781
document "meson" and "metis" in HOL specific section of the Isar ref manual
blanchet
parents:
43270
diff
changeset
|
1924 |
*} |
30169 | 1925 |
|
46280 | 1926 |
|
50130 | 1927 |
section {* Algebraic reasoning via Gr\"obner bases *} |
1928 |
||
1929 |
text {* |
|
1930 |
\begin{matharray}{rcl} |
|
1931 |
@{method_def (HOL) "algebra"} & : & @{text method} \\ |
|
1932 |
@{attribute_def (HOL) algebra} & : & @{text attribute} \\ |
|
1933 |
\end{matharray} |
|
1934 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1935 |
@{rail \<open> |
50130 | 1936 |
@@{method (HOL) algebra} |
1937 |
('add' ':' @{syntax thmrefs})? |
|
1938 |
('del' ':' @{syntax thmrefs})? |
|
1939 |
; |
|
1940 |
@@{attribute (HOL) algebra} (() | 'add' | 'del') |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1941 |
\<close>} |
50130 | 1942 |
|
1943 |
\begin{description} |
|
1944 |
||
1945 |
\item @{method (HOL) algebra} performs algebraic reasoning via |
|
1946 |
Gr\"obner bases, see also \cite{Chaieb-Wenzel:2007} and |
|
1947 |
\cite[\S3.2]{Chaieb-thesis}. The method handles deals with two main |
|
1948 |
classes of problems: |
|
1949 |
||
1950 |
\begin{enumerate} |
|
1951 |
||
1952 |
\item Universal problems over multivariate polynomials in a |
|
1953 |
(semi)-ring/field/idom; the capabilities of the method are augmented |
|
1954 |
according to properties of these structures. For this problem class |
|
1955 |
the method is only complete for algebraically closed fields, since |
|
1956 |
the underlying method is based on Hilbert's Nullstellensatz, where |
|
1957 |
the equivalence only holds for algebraically closed fields. |
|
1958 |
||
1959 |
The problems can contain equations @{text "p = 0"} or inequations |
|
1960 |
@{text "q \<noteq> 0"} anywhere within a universal problem statement. |
|
1961 |
||
1962 |
\item All-exists problems of the following restricted (but useful) |
|
1963 |
form: |
|
1964 |
||
1965 |
@{text [display] "\<forall>x\<^sub>1 \<dots> x\<^sub>n. |
|
1966 |
e\<^sub>1(x\<^sub>1, \<dots>, x\<^sub>n) = 0 \<and> \<dots> \<and> e\<^sub>m(x\<^sub>1, \<dots>, x\<^sub>n) = 0 \<longrightarrow> |
|
1967 |
(\<exists>y\<^sub>1 \<dots> y\<^sub>k. |
|
1968 |
p\<^sub>1\<^sub>1(x\<^sub>1, \<dots> ,x\<^sub>n) * y\<^sub>1 + \<dots> + p\<^sub>1\<^sub>k(x\<^sub>1, \<dots>, x\<^sub>n) * y\<^sub>k = 0 \<and> |
|
1969 |
\<dots> \<and> |
|
1970 |
p\<^sub>t\<^sub>1(x\<^sub>1, \<dots>, x\<^sub>n) * y\<^sub>1 + \<dots> + p\<^sub>t\<^sub>k(x\<^sub>1, \<dots>, x\<^sub>n) * y\<^sub>k = 0)"} |
|
1971 |
||
1972 |
Here @{text "e\<^sub>1, \<dots>, e\<^sub>n"} and the @{text "p\<^sub>i\<^sub>j"} are multivariate |
|
1973 |
polynomials only in the variables mentioned as arguments. |
|
1974 |
||
1975 |
\end{enumerate} |
|
1976 |
||
1977 |
The proof method is preceded by a simplification step, which may be |
|
1978 |
modified by using the form @{text "(algebra add: ths\<^sub>1 del: ths\<^sub>2)"}. |
|
1979 |
This acts like declarations for the Simplifier |
|
1980 |
(\secref{sec:simplifier}) on a private simpset for this tool. |
|
1981 |
||
1982 |
\item @{attribute algebra} (as attribute) manages the default |
|
1983 |
collection of pre-simplification rules of the above proof method. |
|
1984 |
||
1985 |
\end{description} |
|
1986 |
*} |
|
1987 |
||
1988 |
||
1989 |
subsubsection {* Example *} |
|
1990 |
||
1991 |
text {* The subsequent example is from geometry: collinearity is |
|
1992 |
invariant by rotation. *} |
|
1993 |
||
1994 |
type_synonym point = "int \<times> int" |
|
1995 |
||
1996 |
fun collinear :: "point \<Rightarrow> point \<Rightarrow> point \<Rightarrow> bool" where |
|
1997 |
"collinear (Ax, Ay) (Bx, By) (Cx, Cy) \<longleftrightarrow> |
|
1998 |
(Ax - Bx) * (By - Cy) = (Ay - By) * (Bx - Cx)" |
|
1999 |
||
2000 |
lemma collinear_inv_rotation: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52895
diff
changeset
|
2001 |
assumes "collinear (Ax, Ay) (Bx, By) (Cx, Cy)" and "c\<^sup>2 + s\<^sup>2 = 1" |
50130 | 2002 |
shows "collinear (Ax * c - Ay * s, Ay * c + Ax * s) |
2003 |
(Bx * c - By * s, By * c + Bx * s) (Cx * c - Cy * s, Cy * c + Cx * s)" |
|
2004 |
using assms by (algebra add: collinear.simps) |
|
2005 |
||
2006 |
text {* |
|
53982 | 2007 |
See also @{file "~~/src/HOL/ex/Groebner_Examples.thy"}. |
50130 | 2008 |
*} |
2009 |
||
2010 |
||
30171 | 2011 |
section {* Coherent Logic *} |
2012 |
||
2013 |
text {* |
|
2014 |
\begin{matharray}{rcl} |
|
2015 |
@{method_def (HOL) "coherent"} & : & @{text method} \\ |
|
2016 |
\end{matharray} |
|
2017 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2018 |
@{rail \<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2019 |
@@{method (HOL) coherent} @{syntax thmrefs}? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2020 |
\<close>} |
30171 | 2021 |
|
46280 | 2022 |
\begin{description} |
2023 |
||
2024 |
\item @{method (HOL) coherent} solves problems of \emph{Coherent |
|
2025 |
Logic} \cite{Bezem-Coquand:2005}, which covers applications in |
|
2026 |
confluence theory, lattice theory and projective geometry. See |
|
2027 |
@{file "~~/src/HOL/ex/Coherent.thy"} for some examples. |
|
2028 |
||
2029 |
\end{description} |
|
30171 | 2030 |
*} |
2031 |
||
2032 |
||
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2033 |
section {* Proving propositions *} |
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2034 |
|
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2035 |
text {* |
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2036 |
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
|
2037 |
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
|
2038 |
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
|
2039 |
|
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2040 |
\begin{matharray}{rcl} |
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2041 |
@{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
|
2042 |
@{command_def (HOL) "try"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\ |
46641 | 2043 |
@{command_def (HOL) "try0"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\ |
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2044 |
@{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
|
2045 |
@{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
|
2046 |
\end{matharray} |
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2047 |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2048 |
@{rail \<open> |
43040 | 2049 |
@@{command (HOL) try} |
2050 |
; |
|
2051 |
||
46641 | 2052 |
@@{command (HOL) try0} ( ( ( 'simp' | 'intro' | 'elim' | 'dest' ) ':' @{syntax thmrefs} ) + ) ? |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2053 |
@{syntax nat}? |
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2054 |
; |
43040 | 2055 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2056 |
@@{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
|
2057 |
; |
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2058 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2059 |
@@{command (HOL) sledgehammer_params} ( ( '[' args ']' ) ? ) |
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2060 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2061 |
args: ( @{syntax name} '=' value + ',' ) |
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2062 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2063 |
facts: '(' ( ( ( ( 'add' | 'del' ) ':' ) ? @{syntax thmrefs} ) + ) ? ')' |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2064 |
\<close>} % FIXME check args "value" |
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2065 |
|
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2066 |
\begin{description} |
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2067 |
|
46283 | 2068 |
\item @{command (HOL) "solve_direct"} checks whether the current |
2069 |
subgoals can be solved directly by an existing theorem. Duplicate |
|
2070 |
lemmas can be detected in this way. |
|
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2071 |
|
46641 | 2072 |
\item @{command (HOL) "try0"} attempts to prove a subgoal |
46283 | 2073 |
using a combination of standard proof methods (@{method auto}, |
2074 |
@{method simp}, @{method blast}, etc.). Additional facts supplied |
|
2075 |
via @{text "simp:"}, @{text "intro:"}, @{text "elim:"}, and @{text |
|
2076 |
"dest:"} are passed to the appropriate proof methods. |
|
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2077 |
|
43914
64819f353c53
updating documentation about quickcheck; adding information about try
bulwahn
parents:
43578
diff
changeset
|
2078 |
\item @{command (HOL) "try"} attempts to prove or disprove a subgoal |
46283 | 2079 |
using a combination of provers and disprovers (@{command (HOL) |
2080 |
"solve_direct"}, @{command (HOL) "quickcheck"}, @{command (HOL) |
|
46641 | 2081 |
"try0"}, @{command (HOL) "sledgehammer"}, @{command (HOL) |
46283 | 2082 |
"nitpick"}). |
43914
64819f353c53
updating documentation about quickcheck; adding information about try
bulwahn
parents:
43578
diff
changeset
|
2083 |
|
46283 | 2084 |
\item @{command (HOL) "sledgehammer"} attempts to prove a subgoal |
2085 |
using external automatic provers (resolution provers and SMT |
|
2086 |
solvers). See the Sledgehammer manual \cite{isabelle-sledgehammer} |
|
2087 |
for details. |
|
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2088 |
|
46283 | 2089 |
\item @{command (HOL) "sledgehammer_params"} changes @{command (HOL) |
2090 |
"sledgehammer"} configuration options persistently. |
|
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2091 |
|
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2092 |
\end{description} |
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2093 |
*} |
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2094 |
|
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2095 |
|
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2096 |
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
|
2097 |
|
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2098 |
text {* |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2099 |
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
|
2100 |
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
|
2101 |
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
|
2102 |
|
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2103 |
\begin{matharray}{rcl} |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2104 |
@{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
|
2105 |
@{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
|
2106 |
@{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
|
2107 |
@{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
|
2108 |
@{command_def (HOL) "quickcheck_params"} & : & @{text "theory \<rightarrow> theory"} \\ |
45943
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
bulwahn
parents:
45839
diff
changeset
|
2109 |
@{command_def (HOL) "nitpick_params"} & : & @{text "theory \<rightarrow> theory"} \\ |
46592
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2110 |
@{command_def (HOL) "quickcheck_generator"} & : & @{text "theory \<rightarrow> theory"} \\ |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2111 |
@{command_def (HOL) "find_unused_assms"} & : & @{text "context \<rightarrow>"} |
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2112 |
\end{matharray} |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2113 |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2114 |
@{rail \<open> |
58100
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents:
57829
diff
changeset
|
2115 |
@@{command (HOL) value} ( '[' @{syntax 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
|
2116 |
; |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2117 |
|
45409
5abb0e738b00
adding some documentation about the values command to the isar reference
bulwahn
parents:
45408
diff
changeset
|
2118 |
@@{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
|
2119 |
; |
5abb0e738b00
adding some documentation about the values command to the isar reference
bulwahn
parents:
45408
diff
changeset
|
2120 |
|
49993
80402e0e78e3
removed "refute" command from Isar manual, now that it has been moved outside "Main"
blanchet
parents:
49836
diff
changeset
|
2121 |
(@@{command (HOL) quickcheck} | @@{command (HOL) nitpick}) |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2122 |
( '[' args ']' )? @{syntax nat}? |
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2123 |
; |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2124 |
|
49993
80402e0e78e3
removed "refute" command from Isar manual, now that it has been moved outside "Main"
blanchet
parents:
49836
diff
changeset
|
2125 |
(@@{command (HOL) quickcheck_params} | |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2126 |
@@{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
|
2127 |
; |
46592
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2128 |
|
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2129 |
@@{command (HOL) quickcheck_generator} @{syntax nameref} \<newline> |
45943
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
bulwahn
parents:
45839
diff
changeset
|
2130 |
'operations:' ( @{syntax term} +) |
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
bulwahn
parents:
45839
diff
changeset
|
2131 |
; |
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2132 |
|
46628 | 2133 |
@@{command (HOL) find_unused_assms} @{syntax name}? |
46592
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2134 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2135 |
modes: '(' (@{syntax name} +) ')' |
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2136 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2137 |
args: ( @{syntax name} '=' value + ',' ) |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2138 |
\<close>} % FIXME check "value" |
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2139 |
|
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2140 |
\begin{description} |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2141 |
|
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2142 |
\item @{command (HOL) "value"}~@{text t} evaluates and prints a |
46283 | 2143 |
term; optionally @{text modes} can be specified, which are appended |
2144 |
to the current print mode; see \secref{sec:print-modes}. |
|
56927 | 2145 |
Evaluation is tried first using ML, falling |
2146 |
back to normalization by evaluation if this fails. |
|
58100
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents:
57829
diff
changeset
|
2147 |
Alternatively a specific evaluator can be selected using square |
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents:
57829
diff
changeset
|
2148 |
brackets; typical evaluators use the current set of code equations |
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents:
57829
diff
changeset
|
2149 |
to normalize and include @{text simp} for fully symbolic evaluation |
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents:
57829
diff
changeset
|
2150 |
using the simplifier, @{text nbe} for \emph{normalization by |
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents:
57829
diff
changeset
|
2151 |
evaluation} and \emph{code} for code generation in SML. |
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2152 |
|
46283 | 2153 |
\item @{command (HOL) "values"}~@{text t} enumerates a set |
2154 |
comprehension by evaluation and prints its values up to the given |
|
2155 |
number of solutions; optionally @{text modes} can be specified, |
|
2156 |
which are appended to the current print mode; see |
|
2157 |
\secref{sec:print-modes}. |
|
45409
5abb0e738b00
adding some documentation about the values command to the isar reference
bulwahn
parents:
45408
diff
changeset
|
2158 |
|
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2159 |
\item @{command (HOL) "quickcheck"} tests the current goal for |
46283 | 2160 |
counterexamples using a series of assignments for its free |
2161 |
variables; by default the first subgoal is tested, an other can be |
|
2162 |
selected explicitly using an optional goal index. Assignments can |
|
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
2163 |
be chosen exhausting the search space up to a given size, or using a |
46283 | 2164 |
fixed number of random assignments in the search space, or exploring |
2165 |
the search space symbolically using narrowing. By default, |
|
2166 |
quickcheck uses exhaustive testing. A number of configuration |
|
2167 |
options are supported for @{command (HOL) "quickcheck"}, notably: |
|
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2168 |
|
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2169 |
\begin{description} |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2170 |
|
43914
64819f353c53
updating documentation about quickcheck; adding information about try
bulwahn
parents:
43578
diff
changeset
|
2171 |
\item[@{text tester}] specifies which testing approach to apply. |
46283 | 2172 |
There are three testers, @{text exhaustive}, @{text random}, and |
2173 |
@{text narrowing}. An unknown configuration option is treated as |
|
2174 |
an argument to tester, making @{text "tester ="} optional. When |
|
2175 |
multiple testers are given, these are applied in parallel. If no |
|
2176 |
tester is specified, quickcheck uses the testers that are set |
|
2177 |
active, i.e., configurations @{attribute |
|
2178 |
quickcheck_exhaustive_active}, @{attribute |
|
2179 |
quickcheck_random_active}, @{attribute |
|
2180 |
quickcheck_narrowing_active} are set to true. |
|
2181 |
||
40254 | 2182 |
\item[@{text size}] specifies the maximum size of the search space |
2183 |
for assignment values. |
|
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2184 |
|
45758
6210c350d88b
documenting the genuine_only option in quickcheck;
bulwahn
parents:
45701
diff
changeset
|
2185 |
\item[@{text genuine_only}] sets quickcheck only to return genuine |
46283 | 2186 |
counterexample, but not potentially spurious counterexamples due |
2187 |
to underspecified functions. |
|
46498
2754784e9153
adding documentation for abort_potential option in quickcheck
bulwahn
parents:
46494
diff
changeset
|
2188 |
|
2754784e9153
adding documentation for abort_potential option in quickcheck
bulwahn
parents:
46494
diff
changeset
|
2189 |
\item[@{text abort_potential}] sets quickcheck to abort once it |
2754784e9153
adding documentation for abort_potential option in quickcheck
bulwahn
parents:
46494
diff
changeset
|
2190 |
found a potentially spurious counterexample and to not continue |
2754784e9153
adding documentation for abort_potential option in quickcheck
bulwahn
parents:
46494
diff
changeset
|
2191 |
to search for a further genuine counterexample. |
2754784e9153
adding documentation for abort_potential option in quickcheck
bulwahn
parents:
46494
diff
changeset
|
2192 |
For this option to be effective, the @{text genuine_only} option |
2754784e9153
adding documentation for abort_potential option in quickcheck
bulwahn
parents:
46494
diff
changeset
|
2193 |
must be set to false. |
47859 | 2194 |
|
42092
f07b373f25d3
adding documentation about the eval option in quickcheck
bulwahn
parents:
41846
diff
changeset
|
2195 |
\item[@{text eval}] takes a term or a list of terms and evaluates |
46283 | 2196 |
these terms under the variable assignment found by quickcheck. |
48159 | 2197 |
This option is currently only supported by the default |
2198 |
(exhaustive) tester. |
|
42123 | 2199 |
|
40254 | 2200 |
\item[@{text iterations}] sets how many sets of assignments are |
2201 |
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
|
2202 |
|
40254 | 2203 |
\item[@{text no_assms}] specifies whether assumptions in |
2204 |
structured proofs should be ignored. |
|
35331 | 2205 |
|
47349
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
bulwahn
parents:
46641
diff
changeset
|
2206 |
\item[@{text locale}] specifies how to process conjectures in |
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
bulwahn
parents:
46641
diff
changeset
|
2207 |
a locale context, i.e., they can be interpreted or expanded. |
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
bulwahn
parents:
46641
diff
changeset
|
2208 |
The option is a whitespace-separated list of the two words |
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
bulwahn
parents:
46641
diff
changeset
|
2209 |
@{text interpret} and @{text expand}. The list determines the |
47859 | 2210 |
order they are employed. The default setting is to first use |
47349
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
bulwahn
parents:
46641
diff
changeset
|
2211 |
interpretations and then test the expanded conjecture. |
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
bulwahn
parents:
46641
diff
changeset
|
2212 |
The option is only provided as attribute declaration, but not |
47859 | 2213 |
as parameter to the command. |
47349
803729c9fd4d
documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
bulwahn
parents:
46641
diff
changeset
|
2214 |
|
40254 | 2215 |
\item[@{text timeout}] sets the time limit in seconds. |
40245
59f011c1877a
updating documentation on quickcheck in the Isar reference
bulwahn
parents:
40171
diff
changeset
|
2216 |
|
40254 | 2217 |
\item[@{text default_type}] sets the type(s) generally used to |
2218 |
instantiate type variables. |
|
40245
59f011c1877a
updating documentation on quickcheck in the Isar reference
bulwahn
parents:
40171
diff
changeset
|
2219 |
|
40254 | 2220 |
\item[@{text report}] if set quickcheck reports how many tests |
2221 |
fulfilled the preconditions. |
|
40245
59f011c1877a
updating documentation on quickcheck in the Isar reference
bulwahn
parents:
40171
diff
changeset
|
2222 |
|
46592
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2223 |
\item[@{text use_subtype}] if set quickcheck automatically lifts |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2224 |
conjectures to registered subtypes if possible, and tests the |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2225 |
lifted conjecture. |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2226 |
|
45766
46046d8e9659
updating documentation about quiet and verbose options in quickcheck
bulwahn
parents:
45758
diff
changeset
|
2227 |
\item[@{text quiet}] if set quickcheck does not output anything |
46046d8e9659
updating documentation about quiet and verbose options in quickcheck
bulwahn
parents:
45758
diff
changeset
|
2228 |
while testing. |
47859 | 2229 |
|
46283 | 2230 |
\item[@{text verbose}] if set quickcheck informs about the current |
2231 |
size and cardinality while testing. |
|
40245
59f011c1877a
updating documentation on quickcheck in the Isar reference
bulwahn
parents:
40171
diff
changeset
|
2232 |
|
40254 | 2233 |
\item[@{text expect}] can be used to check if the user's |
2234 |
expectation was met (@{text no_expectation}, @{text |
|
2235 |
no_counterexample}, or @{text counterexample}). |
|
40245
59f011c1877a
updating documentation on quickcheck in the Isar reference
bulwahn
parents:
40171
diff
changeset
|
2236 |
|
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2237 |
\end{description} |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2238 |
|
46283 | 2239 |
These option can be given within square brackets. |
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2240 |
|
56363
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2241 |
Using the following type classes, the testers generate values and convert |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2242 |
them back into Isabelle terms for displaying counterexamples. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2243 |
\begin{description} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2244 |
\item[@{text exhaustive}] The parameters of the type classes @{class exhaustive} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2245 |
and @{class full_exhaustive} implement the testing. They take a |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2246 |
testing function as a parameter, which takes a value of type @{typ "'a"} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2247 |
and optionally produces a counterexample, and a size parameter for the test values. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2248 |
In @{class full_exhaustive}, the testing function parameter additionally |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2249 |
expects a lazy term reconstruction in the type @{typ Code_Evaluation.term} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2250 |
of the tested value. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2251 |
|
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2252 |
The canonical implementation for @{text exhaustive} testers calls the given |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2253 |
testing function on all values up to the given size and stops as soon |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2254 |
as a counterexample is found. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2255 |
|
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2256 |
\item[@{text random}] The operation @{const Quickcheck_Random.random} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2257 |
of the type class @{class random} generates a pseudo-random |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2258 |
value of the given size and a lazy term reconstruction of the value |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2259 |
in the type @{typ Code_Evaluation.term}. A pseudo-randomness generator |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2260 |
is defined in theory @{theory Random}. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2261 |
|
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2262 |
\item[@{text narrowing}] implements Haskell's Lazy Smallcheck~\cite{runciman-naylor-lindblad} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2263 |
using the type classes @{class narrowing} and @{class partial_term_of}. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2264 |
Variables in the current goal are initially represented as symbolic variables. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2265 |
If the execution of the goal tries to evaluate one of them, the test engine |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2266 |
replaces it with refinements provided by @{const narrowing}. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2267 |
Narrowing views every value as a sum-of-products which is expressed using the operations |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2268 |
@{const Quickcheck_Narrowing.cons} (embedding a value), |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2269 |
@{const Quickcheck_Narrowing.apply} (product) and @{const Quickcheck_Narrowing.sum} (sum). |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2270 |
The refinement should enable further evaluation of the goal. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2271 |
|
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2272 |
For example, @{const narrowing} for the list type @{typ "'a :: narrowing list"} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2273 |
can be recursively defined as |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2274 |
@{term "Quickcheck_Narrowing.sum (Quickcheck_Narrowing.cons []) |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2275 |
(Quickcheck_Narrowing.apply |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2276 |
(Quickcheck_Narrowing.apply |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2277 |
(Quickcheck_Narrowing.cons (op #)) |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2278 |
narrowing) |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2279 |
narrowing)"}. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2280 |
If a symbolic variable of type @{typ "_ list"} is evaluated, it is replaced by (i)~the empty |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2281 |
list @{term "[]"} and (ii)~by a non-empty list whose head and tail can then be recursively |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2282 |
refined if needed. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2283 |
|
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2284 |
To reconstruct counterexamples, the operation @{const partial_term_of} transforms |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2285 |
@{text narrowing}'s deep representation of terms to the type @{typ Code_Evaluation.term}. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2286 |
The deep representation models symbolic variables as |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2287 |
@{const Quickcheck_Narrowing.Narrowing_variable}, which are normally converted to |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2288 |
@{const Code_Evaluation.Free}, and refined values as |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2289 |
@{term "Quickcheck_Narrowing.Narrowing_constructor i args"}, where @{term "i :: integer"} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2290 |
denotes the index in the sum of refinements. In the above example for lists, |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2291 |
@{term "0"} corresponds to @{term "[]"} and @{term "1"} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2292 |
to @{term "op #"}. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2293 |
|
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2294 |
The command @{command (HOL) "code_datatype"} sets up @{const partial_term_of} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2295 |
such that the @{term "i"}-th refinement is interpreted as the @{term "i"}-th constructor, |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2296 |
but it does not ensures consistency with @{const narrowing}. |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2297 |
\end{description} |
89e0264adf79
document value generation for quickcheck's testers
Andreas Lochbihler
parents:
56270
diff
changeset
|
2298 |
|
46283 | 2299 |
\item @{command (HOL) "quickcheck_params"} changes @{command (HOL) |
2300 |
"quickcheck"} configuration options persistently. |
|
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2301 |
|
45943
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
bulwahn
parents:
45839
diff
changeset
|
2302 |
\item @{command (HOL) "quickcheck_generator"} creates random and |
46283 | 2303 |
exhaustive value generators for a given type and operations. It |
2304 |
generates values by using the operations as if they were |
|
2305 |
constructors of that type. |
|
45943
8c4a5e664fbc
adding documentation about the quickcheck_generator command in the IsarRef
bulwahn
parents:
45839
diff
changeset
|
2306 |
|
46283 | 2307 |
\item @{command (HOL) "nitpick"} tests the current goal for |
2308 |
counterexamples using a reduction to first-order relational |
|
2309 |
logic. See the Nitpick manual \cite{isabelle-nitpick} for details. |
|
42215
de9d43c427ae
document "nitpick(_params)", "refute(_params)", "try", "sledgehammer(_params)", and "solve_direct"
blanchet
parents:
42123
diff
changeset
|
2310 |
|
46283 | 2311 |
\item @{command (HOL) "nitpick_params"} changes @{command (HOL) |
2312 |
"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
|
2313 |
|
46592
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2314 |
\item @{command (HOL) "find_unused_assms"} finds potentially superfluous |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2315 |
assumptions in theorems using quickcheck. |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2316 |
It takes the theory name to be checked for superfluous assumptions as |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2317 |
optional argument. If not provided, it checks the current theory. |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2318 |
Options to the internal quickcheck invocations can be changed with |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2319 |
common configuration declarations. |
d5d49bd4a7b4
adding documentation about find_unused_assms command and use_subtype option in the IsarRef
bulwahn
parents:
46498
diff
changeset
|
2320 |
|
31912
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2321 |
\end{description} |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2322 |
*} |
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2323 |
|
f5bd306f5e9d
more friendly wrt. PGs interpretation of compound *); added dedicated section on value and quickcheck
haftmann
parents:
31254
diff
changeset
|
2324 |
|
28752 | 2325 |
section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *} |
26849 | 2326 |
|
2327 |
text {* |
|
27123
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents:
27103
diff
changeset
|
2328 |
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
|
2329 |
induction in unstructured tactic scripts; see also |
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents:
27103
diff
changeset
|
2330 |
\secref{sec:cases-induct} for proper Isar versions of similar ideas. |
26849 | 2331 |
|
2332 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
2333 |
@{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
|
2334 |
@{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
|
2335 |
@{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
|
2336 |
@{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
26849 | 2337 |
\end{matharray} |
2338 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2339 |
@{rail \<open> |
42705 | 2340 |
@@{method (HOL) case_tac} @{syntax goal_spec}? @{syntax term} rule? |
26849 | 2341 |
; |
42705 | 2342 |
@@{method (HOL) induct_tac} @{syntax goal_spec}? (@{syntax insts} * @'and') rule? |
26849 | 2343 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2344 |
@@{method (HOL) ind_cases} (@{syntax prop}+) (@'for' (@{syntax name}+))? |
26849 | 2345 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2346 |
@@{command (HOL) inductive_cases} (@{syntax thmdecl}? (@{syntax prop}+) + @'and') |
26849 | 2347 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2348 |
rule: 'rule' ':' @{syntax thmref} |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2349 |
\<close>} |
26849 | 2350 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
2351 |
\begin{description} |
26849 | 2352 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
2353 |
\item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
2354 |
to reason about inductive types. Rules are selected according to |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
2355 |
the declarations by the @{attribute cases} and @{attribute induct} |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
2356 |
attributes, cf.\ \secref{sec:cases-induct}. The @{command (HOL) |
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58100
diff
changeset
|
2357 |
datatype_new} package already takes care of this. |
27123
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents:
27103
diff
changeset
|
2358 |
|
11fcdd5897dd
case_tac/induct_tac: use same declarations as cases/induct;
wenzelm
parents:
27103
diff
changeset
|
2359 |
These unstructured tactics feature both goal addressing and dynamic |
26849 | 2360 |
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
|
2361 |
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
|
2362 |
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
|
2363 |
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
|
2364 |
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
|
2365 |
being addressed. |
42123 | 2366 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
2367 |
\item @{method (HOL) ind_cases} and @{command (HOL) |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
2368 |
"inductive_cases"} provide an interface to the internal @{ML_text |
26860 | 2369 |
mk_cases} operation. Rules are simplified in an unrestricted |
2370 |
forward manner. |
|
26849 | 2371 |
|
2372 |
While @{method (HOL) ind_cases} is a proof method to apply the |
|
2373 |
result immediately as elimination rules, @{command (HOL) |
|
2374 |
"inductive_cases"} provides case split theorems at the theory level |
|
2375 |
for later use. The @{keyword "for"} argument of the @{method (HOL) |
|
2376 |
ind_cases} method allows to specify a list of variables that should |
|
2377 |
be generalized before applying the resulting rule. |
|
2378 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28752
diff
changeset
|
2379 |
\end{description} |
26849 | 2380 |
*} |
2381 |
||
2382 |
||
50109 | 2383 |
chapter {* Executable code *} |
26849 | 2384 |
|
42627
8749742785b8
moved material about old codegen to isar-ref manual;
wenzelm
parents:
42626
diff
changeset
|
2385 |
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
|
2386 |
specifications. In principle, execution could be simulated by |
8749742785b8
moved material about old codegen to isar-ref manual;
wenzelm
parents:
42626
diff
changeset
|
2387 |
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
|
2388 |
simplification. Unfortunately, this approach is rather inefficient. |
8749742785b8
moved material about old codegen to isar-ref manual;
wenzelm
parents:
42626
diff
changeset
|
2389 |
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
|
2390 |
them into a functional programming language such as ML. |
26849 | 2391 |
|
45192 | 2392 |
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
|
2393 |
from executable specifications. Isabelle/HOL instantiates these |
45192 | 2394 |
mechanisms in a way that is amenable to end-user applications. Code |
2395 |
can be generated for functional programs (including overloading |
|
2396 |
using type classes) targeting SML \cite{SML}, OCaml \cite{OCaml}, |
|
2397 |
Haskell \cite{haskell-revised-report} and Scala |
|
42627
8749742785b8
moved material about old codegen to isar-ref manual;
wenzelm
parents:
42626
diff
changeset
|
2398 |
\cite{scala-overview-tech-report}. Conceptually, code generation is |
8749742785b8
moved material about old codegen to isar-ref manual;
wenzelm
parents:
42626
diff
changeset
|
2399 |
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
|
2400 |
\emph{translation} into an abstract executable view and |
8749742785b8
moved material about old codegen to isar-ref manual;
wenzelm
parents:
42626
diff
changeset
|
2401 |
\emph{serialization} to a specific \emph{target language}. |
8749742785b8
moved material about old codegen to isar-ref manual;
wenzelm
parents:
42626
diff
changeset
|
2402 |
Inductive specifications can be executed using the predicate |
8749742785b8
moved material about old codegen to isar-ref manual;
wenzelm
parents:
42626
diff
changeset
|
2403 |
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
|
2404 |
an introduction. |
37422 | 2405 |
|
2406 |
\begin{matharray}{rcl} |
|
2407 |
@{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
|
2408 |
@{attribute_def (HOL) code} & : & @{text attribute} \\ |
|
2409 |
@{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
2410 |
@{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
|
2411 |
@{attribute_def (HOL) code_unfold} & : & @{text attribute} \\ |
37422 | 2412 |
@{attribute_def (HOL) code_post} & : & @{text attribute} \\ |
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45944
diff
changeset
|
2413 |
@{attribute_def (HOL) code_abbrev} & : & @{text attribute} \\ |
37422 | 2414 |
@{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
2415 |
@{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
|
2416 |
@{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
|
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2417 |
@{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\ |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2418 |
@{command_def (HOL) "code_printing"} & : & @{text "theory \<rightarrow> theory"} \\ |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2419 |
@{command_def (HOL) "code_identifier"} & : & @{text "theory \<rightarrow> theory"} \\ |
37422 | 2420 |
@{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\ |
45408
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
bulwahn
parents:
45232
diff
changeset
|
2421 |
@{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
|
2422 |
@{command_def (HOL) "code_pred"} & : & @{text "theory \<rightarrow> proof(prove)"} |
37422 | 2423 |
\end{matharray} |
2424 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2425 |
@{rail \<open> |
55686
e99ed112d303
NEWS and documentation, including correction of long-overseen "*"
haftmann
parents:
55677
diff
changeset
|
2426 |
@@{command (HOL) export_code} ( @'open' ) ? ( constexpr + ) \<newline> |
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2427 |
( ( @'in' target ( @'module_name' @{syntax string} ) ? \<newline> |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52378
diff
changeset
|
2428 |
( @'file' @{syntax string} ) ? ( '(' args ')' ) ?) + ) ? |
37422 | 2429 |
; |
2430 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2431 |
const: @{syntax term} |
37422 | 2432 |
; |
2433 |
||
40711
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
haftmann
parents:
40709
diff
changeset
|
2434 |
constexpr: ( const | 'name._' | '_' ) |
37422 | 2435 |
; |
2436 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2437 |
typeconstructor: @{syntax nameref} |
37422 | 2438 |
; |
2439 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2440 |
class: @{syntax nameref} |
37422 | 2441 |
; |
2442 |
||
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2443 |
target: 'SML' | 'OCaml' | 'Haskell' | 'Scala' | 'Eval' |
37422 | 2444 |
; |
2445 |
||
54890
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
54338
diff
changeset
|
2446 |
@@{attribute (HOL) code} ( 'del' | 'equation' | 'abstype' | 'abstract' |
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
54338
diff
changeset
|
2447 |
| 'drop:' ( const + ) | 'abort:' ( const + ) )? |
37422 | 2448 |
; |
2449 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2450 |
@@{command (HOL) code_datatype} ( const + ) |
37422 | 2451 |
; |
2452 |
||
45232
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
bulwahn
parents:
45192
diff
changeset
|
2453 |
@@{attribute (HOL) code_unfold} ( 'del' ) ? |
37422 | 2454 |
; |
2455 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2456 |
@@{attribute (HOL) code_post} ( 'del' ) ? |
37422 | 2457 |
; |
2458 |
||
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45944
diff
changeset
|
2459 |
@@{attribute (HOL) code_abbrev} |
45232
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
bulwahn
parents:
45192
diff
changeset
|
2460 |
; |
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
bulwahn
parents:
45192
diff
changeset
|
2461 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2462 |
@@{command (HOL) code_thms} ( constexpr + ) ? |
37422 | 2463 |
; |
2464 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2465 |
@@{command (HOL) code_deps} ( constexpr + ) ? |
37422 | 2466 |
; |
2467 |
||
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2468 |
@@{command (HOL) code_reserved} target ( @{syntax string} + ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2469 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2470 |
|
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2471 |
symbol_const: ( @'constant' const ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2472 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2473 |
|
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2474 |
symbol_typeconstructor: ( @'type_constructor' typeconstructor ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2475 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2476 |
|
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2477 |
symbol_class: ( @'type_class' class ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2478 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2479 |
|
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2480 |
symbol_class_relation: ( @'class_relation' class ( '<' | '\<subseteq>' ) class ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2481 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2482 |
|
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2483 |
symbol_class_instance: ( @'class_instance' typeconstructor @'::' class ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2484 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2485 |
|
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2486 |
symbol_module: ( @'code_module' name ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2487 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2488 |
|
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2489 |
syntax: @{syntax string} | ( @'infix' | @'infixl' | @'infixr' ) @{syntax nat} @{syntax string} |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2490 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2491 |
|
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2492 |
printing_const: symbol_const ( '\<rightharpoonup>' | '=>' ) \<newline> |
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2493 |
( '(' target ')' syntax ? + @'and' ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2494 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2495 |
|
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2496 |
printing_typeconstructor: symbol_typeconstructor ( '\<rightharpoonup>' | '=>' ) \<newline> |
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2497 |
( '(' target ')' syntax ? + @'and' ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2498 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2499 |
|
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2500 |
printing_class: symbol_class ( '\<rightharpoonup>' | '=>' ) \<newline> |
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2501 |
( '(' target ')' @{syntax string} ? + @'and' ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2502 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2503 |
|
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2504 |
printing_class_relation: symbol_class_relation ( '\<rightharpoonup>' | '=>' ) \<newline> |
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2505 |
( '(' target ')' @{syntax string} ? + @'and' ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2506 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2507 |
|
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2508 |
printing_class_instance: symbol_class_instance ( '\<rightharpoonup>' | '=>' ) \<newline> |
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2509 |
( '(' target ')' '-' ? + @'and' ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2510 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2511 |
|
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2512 |
printing_module: symbol_module ( '\<rightharpoonup>' | '=>' ) \<newline> |
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2513 |
( '(' target ')' ( @{syntax string} ( @'attach' ( const + ) ) ? ) ? + @'and' ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2514 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2515 |
|
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2516 |
@@{command (HOL) code_printing} ( ( printing_const | printing_typeconstructor |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2517 |
| printing_class | printing_class_relation | printing_class_instance |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2518 |
| printing_module ) + '|' ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2519 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2520 |
|
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2521 |
@@{command (HOL) code_identifier} ( ( symbol_const | symbol_typeconstructor |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2522 |
| symbol_class | symbol_class_relation | symbol_class_instance |
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2523 |
| symbol_module ) ( '\<rightharpoonup>' | '=>' ) \<newline> |
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2524 |
( '(' target ')' @{syntax string} ? + @'and' ) + '|' ) |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2525 |
; |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2526 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2527 |
@@{command (HOL) code_monad} const const target |
37422 | 2528 |
; |
2529 |
||
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2530 |
@@{command (HOL) code_reflect} @{syntax string} \<newline> |
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2531 |
( @'datatypes' ( @{syntax string} '=' ( '_' | ( @{syntax string} + '|' ) + @'and' ) ) ) ? \<newline> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
42215
diff
changeset
|
2532 |
( @'functions' ( @{syntax string} + ) ) ? ( @'file' @{syntax string} ) ? |
37422 | 2533 |
; |
2534 |
||
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2535 |
@@{command (HOL) code_pred} \<newline> ('(' @'modes' ':' modedecl ')')? \<newline> const |
45408
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
bulwahn
parents:
45232
diff
changeset
|
2536 |
; |
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
bulwahn
parents:
45232
diff
changeset
|
2537 |
|
55029
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2538 |
modedecl: (modes | ((const ':' modes) \<newline> |
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents:
54890
diff
changeset
|
2539 |
(@'and' ((const ':' modes @'and') +))?)) |
45408
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
bulwahn
parents:
45232
diff
changeset
|
2540 |
; |
47859 | 2541 |
|
45408
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
bulwahn
parents:
45232
diff
changeset
|
2542 |
modes: mode @'as' const |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
2543 |
\<close>} |
37422 | 2544 |
|
2545 |
\begin{description} |
|
2546 |
||
2547 |
\item @{command (HOL) "export_code"} generates code for a given list |
|
39608 | 2548 |
of constants in the specified target language(s). If no |
2549 |
serialization instruction is given, only abstract code is generated |
|
2550 |
internally. |
|
37422 | 2551 |
|
2552 |
Constants may be specified by giving them literally, referring to |
|
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
2553 |
all executable constants within a certain theory by giving @{text |
55686
e99ed112d303
NEWS and documentation, including correction of long-overseen "*"
haftmann
parents:
55677
diff
changeset
|
2554 |
"name._"}, or referring to \emph{all} executable constants currently |
e99ed112d303
NEWS and documentation, including correction of long-overseen "*"
haftmann
parents:
55677
diff
changeset
|
2555 |
available by giving @{text "_"}. |
e99ed112d303
NEWS and documentation, including correction of long-overseen "*"
haftmann
parents:
55677
diff
changeset
|
2556 |
|
e99ed112d303
NEWS and documentation, including correction of long-overseen "*"
haftmann
parents:
55677
diff
changeset
|
2557 |
By default, exported identifiers are minimized per module. This |
e99ed112d303
NEWS and documentation, including correction of long-overseen "*"
haftmann
parents:
55677
diff
changeset
|
2558 |
can be suppressed by prepending @{keyword "open"} before the list |
e99ed112d303
NEWS and documentation, including correction of long-overseen "*"
haftmann
parents:
55677
diff
changeset
|
2559 |
of contants. |
37422 | 2560 |
|
2561 |
By default, for each involved theory one corresponding name space |
|
52895
a806aa7a5370
some documentation for adhoc overloading;
Christian Sternagel
parents:
52637
diff
changeset
|
2562 |
module is generated. Alternatively, a module name may be specified |
37422 | 2563 |
after the @{keyword "module_name"} keyword; then \emph{all} code is |
2564 |
placed in this module. |
|
2565 |
||
39608 | 2566 |
For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification |
2567 |
refers to a single file; for \emph{Haskell}, it refers to a whole |
|
2568 |
directory, where code is generated in multiple files reflecting the |
|
2569 |
module hierarchy. Omitting the file specification denotes standard |
|
37749 | 2570 |
output. |
37422 | 2571 |
|
55677 | 2572 |
Serializers take an optional list of arguments in parentheses. |
39608 | 2573 |
For \emph{Haskell} a module name prefix may be given using the |
2574 |
``@{text "root:"}'' argument; ``@{text string_classes}'' adds a |
|
2575 |
``@{verbatim "deriving (Read, Show)"}'' clause to each appropriate |
|
2576 |
datatype declaration. |
|
37422 | 2577 |
|
52637
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2578 |
\item @{attribute (HOL) code} declare code equations for code |
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2579 |
generation. Variant @{text "code equation"} declares a conventional |
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2580 |
equation as code equation. Variants @{text "code abstype"} and |
38462
34d3de1254cd
formally document `code abstype` and `code abstract` attributes
haftmann
parents:
37820
diff
changeset
|
2581 |
@{text "code abstract"} declare abstract datatype certificates or |
34d3de1254cd
formally document `code abstype` and `code abstract` attributes
haftmann
parents:
37820
diff
changeset
|
2582 |
code equations on abstract datatype representations respectively. |
52637
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2583 |
Vanilla @{text "code"} falls back to @{text "code equation"} |
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2584 |
or @{text "code abstype"} depending on the syntactic shape |
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2585 |
of the underlying equation. Variant @{text "code del"} |
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2586 |
deselects a code equation for code generation. |
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2587 |
|
54890
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
54338
diff
changeset
|
2588 |
Variants @{text "code drop:"} and @{text "code abort:"} take |
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
54338
diff
changeset
|
2589 |
a list of constant as arguments and drop all code equations declared |
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
54338
diff
changeset
|
2590 |
for them. In the case of {text abort}, these constants then are |
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
54338
diff
changeset
|
2591 |
are not required to have a definition by means of code equations; |
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
54338
diff
changeset
|
2592 |
if needed these are implemented by program abort (exception) instead. |
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
54338
diff
changeset
|
2593 |
|
52637
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2594 |
Usually packages introducing code equations provide a reasonable |
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52476
diff
changeset
|
2595 |
default setup for selection. |
37422 | 2596 |
|
2597 |
\item @{command (HOL) "code_datatype"} specifies a constructor set |
|
2598 |
for a logical type. |
|
2599 |
||
2600 |
\item @{command (HOL) "print_codesetup"} gives an overview on |
|
2601 |
selected code equations and code generator datatypes. |
|
2602 |
||
45232
eb56e1774c26
updating documentation: code_inline -> code_unfold; added documentation about attribute code_unfold_post
bulwahn
parents:
45192
diff
changeset
|
2603 |
\item @{attribute (HOL) code_unfold} declares (or with option |
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45944
diff
changeset
|
2604 |
``@{text "del"}'' removes) theorems which during preprocessing |
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45944
diff
changeset
|
2605 |
are applied as rewrite rules to any code equation or evaluation |
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45944
diff
changeset
|
2606 |
input. |
37422 | 2607 |
|
39608 | 2608 |
\item @{attribute (HOL) code_post} declares (or with option ``@{text |
2609 |
"del"}'' removes) theorems which are applied as rewrite rules to any |
|
2610 |
result of an evaluation. |
|
37422 | 2611 |
|
56929 | 2612 |
\item @{attribute (HOL) code_abbrev} declares (or with option ``@{text |
2613 |
"del"}'' removes) equations which are |
|
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45944
diff
changeset
|
2614 |
applied as rewrite rules to any result of an evaluation and |
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45944
diff
changeset
|
2615 |
symmetrically during preprocessing to any code equation or evaluation |
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45944
diff
changeset
|
2616 |
input. |
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45944
diff
changeset
|
2617 |
|
39608 | 2618 |
\item @{command (HOL) "print_codeproc"} prints the setup of the code |
2619 |
generator preprocessor. |
|
37422 | 2620 |
|
2621 |
\item @{command (HOL) "code_thms"} prints a list of theorems |
|
2622 |
representing the corresponding program containing all given |
|
2623 |
constants after preprocessing. |
|
2624 |
||
2625 |
\item @{command (HOL) "code_deps"} visualizes dependencies of |
|
2626 |
theorems representing the corresponding program containing all given |
|
2627 |
constants after preprocessing. |
|
2628 |
||
2629 |
\item @{command (HOL) "code_reserved"} declares a list of names as |
|
2630 |
reserved for a given target, preventing it to be shadowed by any |
|
2631 |
generated code. |
|
2632 |
||
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2633 |
\item @{command (HOL) "code_printing"} associates a series of symbols |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2634 |
(constants, type constructors, classes, class relations, instances, |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2635 |
module names) with target-specific serializations; omitting a serialization |
55372 | 2636 |
deletes an existing serialization. |
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2637 |
|
37422 | 2638 |
\item @{command (HOL) "code_monad"} provides an auxiliary mechanism |
2639 |
to generate monadic code for Haskell. |
|
2640 |
||
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2641 |
\item @{command (HOL) "code_identifier"} associates a a series of symbols |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2642 |
(constants, type constructors, classes, class relations, instances, |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
50879
diff
changeset
|
2643 |
module names) with target-specific hints how these symbols shall be named. |
55293
42cf5802d36a
code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents:
55112
diff
changeset
|
2644 |
These hints gain precedence over names for symbols with no hints at all. |
42cf5802d36a
code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents:
55112
diff
changeset
|
2645 |
Conflicting hints are subject to name disambiguation. |
42cf5802d36a
code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents:
55112
diff
changeset
|
2646 |
\emph{Warning:} It is at the discretion |
52476 | 2647 |
of the user to ensure that name prefixes of identifiers in compound |
2648 |
statements like type classes or datatypes are still the same. |
|
37422 | 2649 |
|
39608 | 2650 |
\item @{command (HOL) "code_reflect"} without a ``@{text "file"}'' |
2651 |
argument compiles code into the system runtime environment and |
|
2652 |
modifies the code generator setup that future invocations of system |
|
2653 |
runtime code generation referring to one of the ``@{text |
|
47859 | 2654 |
"datatypes"}'' or ``@{text "functions"}'' entities use these |
2655 |
precompiled entities. With a ``@{text "file"}'' argument, the |
|
2656 |
corresponding code is generated into that specified file without |
|
2657 |
modifying the code generator setup. |
|
2658 |
||
2659 |
\item @{command (HOL) "code_pred"} creates code equations for a |
|
2660 |
predicate given a set of introduction rules. Optional mode |
|
2661 |
annotations determine which arguments are supposed to be input or |
|
2662 |
output. If alternative introduction rules are declared, one must |
|
2663 |
prove a corresponding elimination rule. |
|
45408
7156f63ce3c2
adding a minimal documentation about the code_pred command to the isar reference
bulwahn
parents:
45232
diff
changeset
|
2664 |
|
37422 | 2665 |
\end{description} |
42627
8749742785b8
moved material about old codegen to isar-ref manual;
wenzelm
parents:
42626
diff
changeset
|
2666 |
*} |
37422 | 2667 |
|
26840 | 2668 |
end |