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