| author | wenzelm | 
| Fri, 30 Aug 2013 13:46:32 +0200 | |
| changeset 53327 | d0e4c8f73541 | 
| parent 52753 | 1165f78c16d8 | 
| child 58620 | 7435b6a3f72e | 
| permissions | -rw-r--r-- | 
| 37613 | 1  | 
theory Inductive_Predicate  | 
2  | 
imports Setup  | 
|
3  | 
begin  | 
|
4  | 
||
5  | 
(*<*)  | 
|
| 38441 | 6  | 
hide_const %invisible append  | 
| 37613 | 7  | 
|
| 38441 | 8  | 
inductive %invisible append where  | 
| 37613 | 9  | 
"append [] ys ys"  | 
| 38811 | 10  | 
| "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"  | 
| 38441 | 11  | 
|
12  | 
lemma %invisible append: "append xs ys zs = (xs @ ys = zs)"  | 
|
13  | 
by (induct xs arbitrary: ys zs) (auto elim: append.cases intro: append.intros)  | 
|
| 46515 | 14  | 
|
15  | 
lemmas lexordp_def =  | 
|
16  | 
lexordp_def [unfolded lexord_def mem_Collect_eq split]  | 
|
| 37613 | 17  | 
(*>*)  | 
| 38441 | 18  | 
|
19  | 
section {* Inductive Predicates \label{sec:inductive} *}
 | 
|
20  | 
||
| 37613 | 21  | 
text {*
 | 
| 38508 | 22  | 
  The @{text "predicate compiler"} is an extension of the code generator
 | 
| 38441 | 23  | 
which turns inductive specifications into equational ones, from  | 
24  | 
which in turn executable code can be generated. The mechanisms of  | 
|
25  | 
this compiler are described in detail in  | 
|
26  | 
  \cite{Berghofer-Bulwahn-Haftmann:2009:TPHOL}.
 | 
|
27  | 
||
28  | 
  Consider the simple predicate @{const append} given by these two
 | 
|
29  | 
introduction rules:  | 
|
| 37613 | 30  | 
*}  | 
| 38441 | 31  | 
|
32  | 
text %quote {*
 | 
|
33  | 
  @{thm append.intros(1)[of ys]} \\
 | 
|
34  | 
  @{thm append.intros(2)[of xs ys zs x]}
 | 
|
35  | 
*}  | 
|
36  | 
||
| 37613 | 37  | 
text {*
 | 
| 38441 | 38  | 
  \noindent To invoke the compiler, simply use @{command_def "code_pred"}:
 | 
| 37613 | 39  | 
*}  | 
40  | 
||
| 38441 | 41  | 
code_pred %quote append .  | 
42  | 
||
43  | 
text {*
 | 
|
44  | 
  \noindent The @{command "code_pred"} command takes the name of the
 | 
|
45  | 
inductive predicate and then you put a period to discharge a trivial  | 
|
46  | 
correctness proof. The compiler infers possible modes for the  | 
|
47  | 
predicate and produces the derived code equations. Modes annotate  | 
|
48  | 
which (parts of the) arguments are to be taken as input, and which  | 
|
49  | 
  output. Modes are similar to types, but use the notation @{text "i"}
 | 
|
50  | 
  for input and @{text "o"} for output.
 | 
|
51  | 
||
52  | 
  For @{term "append"}, the compiler can infer the following modes:
 | 
|
53  | 
  \begin{itemize}
 | 
|
54  | 
    \item @{text "i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool"}
 | 
|
55  | 
    \item @{text "i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool"}
 | 
|
56  | 
    \item @{text "o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool"}
 | 
|
57  | 
  \end{itemize}
 | 
|
58  | 
  You can compute sets of predicates using @{command_def "values"}:
 | 
|
59  | 
*}  | 
|
60  | 
||
61  | 
values %quote "{zs. append [(1::nat),2,3] [4,5] zs}"
 | 
|
62  | 
||
63  | 
text {* \noindent outputs @{text "{[1, 2, 3, 4, 5]}"}, and *}
 | 
|
64  | 
||
65  | 
values %quote "{(xs, ys). append xs ys [(2::nat),3]}"
 | 
|
66  | 
||
67  | 
text {* \noindent outputs @{text "{([], [2, 3]), ([2], [3]), ([2, 3], [])}"}. *}
 | 
|
| 37613 | 68  | 
|
69  | 
text {*
 | 
|
| 38441 | 70  | 
\noindent If you are only interested in the first elements of the  | 
71  | 
set comprehension (with respect to a depth-first search on the  | 
|
72  | 
  introduction rules), you can pass an argument to @{command "values"}
 | 
|
73  | 
to specify the number of elements you want:  | 
|
74  | 
*}  | 
|
| 37613 | 75  | 
|
| 38441 | 76  | 
values %quote 1 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}"
 | 
77  | 
values %quote 3 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}"
 | 
|
78  | 
||
| 37613 | 79  | 
text {*
 | 
| 38441 | 80  | 
  \noindent The @{command "values"} command can only compute set
 | 
81  | 
comprehensions for which a mode has been inferred.  | 
|
82  | 
||
83  | 
The code equations for a predicate are made available as theorems with  | 
|
84  | 
  the suffix @{text "equation"}, and can be inspected with:
 | 
|
| 37613 | 85  | 
*}  | 
| 38441 | 86  | 
|
87  | 
thm %quote append.equation  | 
|
88  | 
||
89  | 
text {*
 | 
|
90  | 
\noindent More advanced options are described in the following subsections.  | 
|
91  | 
*}  | 
|
92  | 
||
93  | 
subsection {* Alternative names for functions *}
 | 
|
94  | 
||
| 37613 | 95  | 
text {* 
 | 
| 38441 | 96  | 
By default, the functions generated from a predicate are named after  | 
97  | 
  the predicate with the mode mangled into the name (e.g., @{text
 | 
|
98  | 
"append_i_i_o"}). You can specify your own names as follows:  | 
|
| 37613 | 99  | 
*}  | 
| 38441 | 100  | 
|
| 38811 | 101  | 
code_pred %quote (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool as concat,  | 
102  | 
o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as split,  | 
|
103  | 
i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as suffix) append .  | 
|
| 37613 | 104  | 
|
| 38441 | 105  | 
subsection {* Alternative introduction rules *}
 | 
106  | 
||
| 37613 | 107  | 
text {*
 | 
| 38441 | 108  | 
Sometimes the introduction rules of an predicate are not executable  | 
109  | 
because they contain non-executable constants or specific modes  | 
|
110  | 
could not be inferred. It is also possible that the introduction  | 
|
111  | 
rules yield a function that loops forever due to the execution in a  | 
|
112  | 
depth-first search manner. Therefore, you can declare alternative  | 
|
113  | 
  introduction rules for predicates with the attribute @{attribute
 | 
|
114  | 
"code_pred_intro"}. For example, the transitive closure is defined  | 
|
115  | 
by:  | 
|
| 37613 | 116  | 
*}  | 
| 38441 | 117  | 
|
| 37613 | 118  | 
text %quote {*
 | 
| 
39682
 
066e2d4d0de8
avoid fragile tranclp syntax; corrected resolution; corrected typo
 
haftmann 
parents: 
39065 
diff
changeset
 | 
119  | 
  @{lemma [source] "r a b \<Longrightarrow> tranclp r a b" by (fact tranclp.intros(1))}\\
 | 
| 
 
066e2d4d0de8
avoid fragile tranclp syntax; corrected resolution; corrected typo
 
haftmann 
parents: 
39065 
diff
changeset
 | 
120  | 
  @{lemma [source] "tranclp r a b \<Longrightarrow> r b c \<Longrightarrow> tranclp r a c" by (fact tranclp.intros(2))}
 | 
| 37613 | 121  | 
*}  | 
| 38441 | 122  | 
|
| 37613 | 123  | 
text {*
 | 
| 38441 | 124  | 
\noindent These rules do not suit well for executing the transitive  | 
125  | 
  closure with the mode @{text "(i \<Rightarrow> o \<Rightarrow> bool) \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool"}, as
 | 
|
126  | 
the second rule will cause an infinite loop in the recursive call.  | 
|
127  | 
This can be avoided using the following alternative rules which are  | 
|
128  | 
  declared to the predicate compiler by the attribute @{attribute
 | 
|
129  | 
"code_pred_intro"}:  | 
|
| 37613 | 130  | 
*}  | 
| 38441 | 131  | 
|
| 37613 | 132  | 
lemma %quote [code_pred_intro]:  | 
| 
39682
 
066e2d4d0de8
avoid fragile tranclp syntax; corrected resolution; corrected typo
 
haftmann 
parents: 
39065 
diff
changeset
 | 
133  | 
"r a b \<Longrightarrow> tranclp r a b"  | 
| 
 
066e2d4d0de8
avoid fragile tranclp syntax; corrected resolution; corrected typo
 
haftmann 
parents: 
39065 
diff
changeset
 | 
134  | 
"r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"  | 
| 37613 | 135  | 
by auto  | 
| 38441 | 136  | 
|
| 37613 | 137  | 
text {*
 | 
| 38441 | 138  | 
\noindent After declaring all alternative rules for the transitive  | 
139  | 
  closure, you invoke @{command "code_pred"} as usual.  As you have
 | 
|
140  | 
declared alternative rules for the predicate, you are urged to prove  | 
|
141  | 
that these introduction rules are complete, i.e., that you can  | 
|
142  | 
derive an elimination rule for the alternative rules:  | 
|
| 37613 | 143  | 
*}  | 
| 38441 | 144  | 
|
| 37613 | 145  | 
code_pred %quote tranclp  | 
146  | 
proof -  | 
|
147  | 
case tranclp  | 
|
| 
39682
 
066e2d4d0de8
avoid fragile tranclp syntax; corrected resolution; corrected typo
 
haftmann 
parents: 
39065 
diff
changeset
 | 
148  | 
from this converse_tranclpE [OF tranclp.prems] show thesis by metis  | 
| 37613 | 149  | 
qed  | 
| 38441 | 150  | 
|
| 37613 | 151  | 
text {*
 | 
| 38441 | 152  | 
\noindent Alternative rules can also be used for constants that have  | 
153  | 
not been defined inductively. For example, the lexicographic order  | 
|
154  | 
which is defined as:  | 
|
| 37613 | 155  | 
*}  | 
| 38441 | 156  | 
|
157  | 
text %quote {*
 | 
|
| 46515 | 158  | 
  @{thm [display] lexordp_def [of r]}
 | 
| 37613 | 159  | 
*}  | 
| 38441 | 160  | 
|
161  | 
text {*
 | 
|
162  | 
\noindent To make it executable, you can derive the following two  | 
|
163  | 
rules and prove the elimination rule:  | 
|
164  | 
*}  | 
|
165  | 
||
| 37613 | 166  | 
lemma %quote [code_pred_intro]:  | 
| 46515 | 167  | 
"append xs (a # v) ys \<Longrightarrow> lexordp r xs ys"  | 
168  | 
(*<*)unfolding lexordp_def by (auto simp add: append)(*>*)  | 
|
| 37613 | 169  | 
|
170  | 
lemma %quote [code_pred_intro]:  | 
|
| 46515 | 171  | 
"append u (a # v) xs \<Longrightarrow> append u (b # w) ys \<Longrightarrow> r a b  | 
172  | 
\<Longrightarrow> lexordp r xs ys"  | 
|
173  | 
(*<*)unfolding lexordp_def append apply simp  | 
|
| 38441 | 174  | 
apply (rule disjI2) by auto(*>*)  | 
| 37613 | 175  | 
|
| 46515 | 176  | 
code_pred %quote lexordp  | 
| 38441 | 177  | 
(*<*)proof -  | 
| 37613 | 178  | 
fix r xs ys  | 
| 46515 | 179  | 
assume lexord: "lexordp r xs ys"  | 
180  | 
assume 1: "\<And>r' xs' ys' a v. r = r' \<Longrightarrow> xs = xs' \<Longrightarrow> ys = ys'  | 
|
181  | 
\<Longrightarrow> append xs' (a # v) ys' \<Longrightarrow> thesis"  | 
|
182  | 
assume 2: "\<And>r' xs' ys' u a v b w. r = r' \<Longrightarrow> xs = xs' \<Longrightarrow> ys = ys'  | 
|
183  | 
\<Longrightarrow> append u (a # v) xs' \<Longrightarrow> append u (b # w) ys' \<Longrightarrow> r' a b \<Longrightarrow> thesis"  | 
|
| 37613 | 184  | 
  {
 | 
185  | 
assume "\<exists>a v. ys = xs @ a # v"  | 
|
186  | 
from this 1 have thesis  | 
|
| 46515 | 187  | 
by (fastforce simp add: append)  | 
| 37613 | 188  | 
} moreover  | 
189  | 
  {
 | 
|
| 46515 | 190  | 
assume "\<exists>u a b v w. r a b \<and> xs = u @ a # v \<and> ys = u @ b # w"  | 
191  | 
from this 2 have thesis by (fastforce simp add: append)  | 
|
| 37613 | 192  | 
} moreover  | 
193  | 
note lexord  | 
|
194  | 
ultimately show thesis  | 
|
| 46515 | 195  | 
unfolding lexordp_def  | 
196  | 
by fastforce  | 
|
| 38441 | 197  | 
qed(*>*)  | 
198  | 
||
199  | 
||
200  | 
subsection {* Options for values *}
 | 
|
201  | 
||
| 37613 | 202  | 
text {*
 | 
| 38441 | 203  | 
In the presence of higher-order predicates, multiple modes for some  | 
204  | 
predicate could be inferred that are not disambiguated by the  | 
|
205  | 
pattern of the set comprehension. To disambiguate the modes for the  | 
|
206  | 
arguments of a predicate, you can state the modes explicitly in the  | 
|
207  | 
  @{command "values"} command.  Consider the simple predicate @{term
 | 
|
208  | 
"succ"}:  | 
|
| 37613 | 209  | 
*}  | 
| 38441 | 210  | 
|
211  | 
inductive %quote succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where  | 
|
| 37613 | 212  | 
"succ 0 (Suc 0)"  | 
213  | 
| "succ x y \<Longrightarrow> succ (Suc x) (Suc y)"  | 
|
214  | 
||
| 38441 | 215  | 
code_pred %quote succ .  | 
| 37613 | 216  | 
|
217  | 
text {*
 | 
|
| 38441 | 218  | 
  \noindent For this, the predicate compiler can infer modes @{text "o
 | 
219  | 
  \<Rightarrow> o \<Rightarrow> bool"}, @{text "i \<Rightarrow> o \<Rightarrow> bool"}, @{text "o \<Rightarrow> i \<Rightarrow> bool"} and
 | 
|
220  | 
  @{text "i \<Rightarrow> i \<Rightarrow> bool"}.  The invocation of @{command "values"}
 | 
|
221  | 
  @{text "{n. tranclp succ 10 n}"} loops, as multiple modes for the
 | 
|
222  | 
  predicate @{text "succ"} are possible and here the first mode @{text
 | 
|
223  | 
"o \<Rightarrow> o \<Rightarrow> bool"} is chosen. To choose another mode for the argument,  | 
|
224  | 
  you can declare the mode for the argument between the @{command
 | 
|
225  | 
"values"} and the number of elements.  | 
|
| 37613 | 226  | 
*}  | 
| 38441 | 227  | 
|
| 39065 | 228  | 
values %quote [mode: i \<Rightarrow> o \<Rightarrow> bool] 1 "{n. tranclp succ 10 n}" (*FIMXE does not terminate for n\<ge>1*)
 | 
229  | 
values %quote [mode: o \<Rightarrow> i \<Rightarrow> bool] 1 "{n. tranclp succ n 10}"
 | 
|
| 37613 | 230  | 
|
| 38441 | 231  | 
|
232  | 
subsection {* Embedding into functional code within Isabelle/HOL *}
 | 
|
233  | 
||
| 37613 | 234  | 
text {*
 | 
| 38441 | 235  | 
To embed the computation of an inductive predicate into functions  | 
236  | 
that are defined in Isabelle/HOL, you have a number of options:  | 
|
237  | 
||
238  | 
  \begin{itemize}
 | 
|
239  | 
||
240  | 
\item You want to use the first-order predicate with the mode  | 
|
241  | 
where all arguments are input. Then you can use the predicate directly, e.g.  | 
|
242  | 
||
243  | 
      \begin{quote}
 | 
|
244  | 
        @{text "valid_suffix ys zs = "} \\
 | 
|
245  | 
        @{text "(if append [Suc 0, 2] ys zs then Some ys else None)"}
 | 
|
246  | 
      \end{quote}
 | 
|
247  | 
||
248  | 
\item If you know that the execution returns only one value (it is  | 
|
249  | 
      deterministic), then you can use the combinator @{term
 | 
|
250  | 
"Predicate.the"}, e.g., a functional concatenation of lists is  | 
|
251  | 
defined with  | 
|
252  | 
||
253  | 
      \begin{quote}
 | 
|
254  | 
        @{term "functional_concat xs ys = Predicate.the (append_i_i_o xs ys)"}
 | 
|
255  | 
      \end{quote}
 | 
|
256  | 
||
257  | 
Note that if the evaluation does not return a unique value, it  | 
|
258  | 
      raises a run-time error @{term "not_unique"}.
 | 
|
259  | 
||
260  | 
  \end{itemize}
 | 
|
| 37613 | 261  | 
*}  | 
| 38441 | 262  | 
|
263  | 
||
264  | 
subsection {* Further Examples *}
 | 
|
265  | 
||
266  | 
text {*
 | 
|
267  | 
Further examples for compiling inductive predicates can be found in  | 
|
| 52753 | 268  | 
  @{file "~~/src/HOL/Predicate_Compile_Examples/Examples.thy"}.  There are
 | 
| 38441 | 269  | 
also some examples in the Archive of Formal Proofs, notably in the  | 
270  | 
  @{text "POPLmark-deBruijn"} and the @{text "FeatherweightJava"}
 | 
|
271  | 
sessions.  | 
|
| 37613 | 272  | 
*}  | 
| 38441 | 273  | 
|
| 37613 | 274  | 
end  | 
| 46515 | 275  |