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