haftmann@37613: theory Inductive_Predicate haftmann@37613: imports Setup haftmann@37613: begin haftmann@37613: haftmann@37613: (*<*) haftmann@38441: hide_const %invisible append haftmann@37613: haftmann@38441: inductive %invisible append where haftmann@37613: "append [] ys ys" haftmann@38811: | "append xs ys zs \ append (x # xs) ys (x # zs)" haftmann@38441: haftmann@38441: lemma %invisible append: "append xs ys zs = (xs @ ys = zs)" haftmann@38441: by (induct xs arbitrary: ys zs) (auto elim: append.cases intro: append.intros) haftmann@46515: haftmann@46515: lemmas lexordp_def = haftmann@46515: lexordp_def [unfolded lexord_def mem_Collect_eq split] haftmann@37613: (*>*) haftmann@38441: haftmann@59377: section \Inductive Predicates \label{sec:inductive}\ haftmann@38441: haftmann@59377: text \ haftmann@38508: The @{text "predicate compiler"} is an extension of the code generator haftmann@38441: which turns inductive specifications into equational ones, from haftmann@38441: which in turn executable code can be generated. The mechanisms of haftmann@38441: this compiler are described in detail in wenzelm@58620: @{cite "Berghofer-Bulwahn-Haftmann:2009:TPHOL"}. haftmann@38441: haftmann@38441: Consider the simple predicate @{const append} given by these two haftmann@38441: introduction rules: haftmann@59377: \ haftmann@38441: haftmann@59377: text %quote \ haftmann@38441: @{thm append.intros(1)[of ys]} \\ haftmann@38441: @{thm append.intros(2)[of xs ys zs x]} haftmann@59377: \ haftmann@38441: haftmann@59377: text \ haftmann@38441: \noindent To invoke the compiler, simply use @{command_def "code_pred"}: haftmann@59377: \ haftmann@37613: haftmann@38441: code_pred %quote append . haftmann@38441: haftmann@59377: text \ haftmann@38441: \noindent The @{command "code_pred"} command takes the name of the haftmann@38441: inductive predicate and then you put a period to discharge a trivial haftmann@38441: correctness proof. The compiler infers possible modes for the haftmann@38441: predicate and produces the derived code equations. Modes annotate haftmann@38441: which (parts of the) arguments are to be taken as input, and which haftmann@38441: output. Modes are similar to types, but use the notation @{text "i"} haftmann@38441: for input and @{text "o"} for output. haftmann@38441: haftmann@38441: For @{term "append"}, the compiler can infer the following modes: haftmann@38441: \begin{itemize} haftmann@38441: \item @{text "i \ i \ i \ bool"} haftmann@38441: \item @{text "i \ i \ o \ bool"} haftmann@38441: \item @{text "o \ o \ i \ bool"} haftmann@38441: \end{itemize} haftmann@38441: You can compute sets of predicates using @{command_def "values"}: haftmann@59377: \ haftmann@38441: haftmann@38441: values %quote "{zs. append [(1::nat),2,3] [4,5] zs}" haftmann@38441: haftmann@59377: text \\noindent outputs @{text "{[1, 2, 3, 4, 5]}"}, and\ haftmann@38441: haftmann@38441: values %quote "{(xs, ys). append xs ys [(2::nat),3]}" haftmann@38441: haftmann@59377: text \\noindent outputs @{text "{([], [2, 3]), ([2], [3]), ([2, 3], [])}"}.\ haftmann@37613: haftmann@59377: text \ haftmann@38441: \noindent If you are only interested in the first elements of the haftmann@38441: set comprehension (with respect to a depth-first search on the haftmann@38441: introduction rules), you can pass an argument to @{command "values"} haftmann@38441: to specify the number of elements you want: haftmann@59377: \ haftmann@37613: haftmann@38441: values %quote 1 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}" haftmann@38441: values %quote 3 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}" haftmann@38441: haftmann@59377: text \ haftmann@38441: \noindent The @{command "values"} command can only compute set haftmann@38441: comprehensions for which a mode has been inferred. haftmann@38441: haftmann@38441: The code equations for a predicate are made available as theorems with haftmann@38441: the suffix @{text "equation"}, and can be inspected with: haftmann@59377: \ haftmann@38441: haftmann@38441: thm %quote append.equation haftmann@38441: haftmann@59377: text \ haftmann@38441: \noindent More advanced options are described in the following subsections. haftmann@59377: \ haftmann@38441: haftmann@59377: subsection \Alternative names for functions\ haftmann@38441: haftmann@59377: text \ haftmann@38441: By default, the functions generated from a predicate are named after haftmann@38441: the predicate with the mode mangled into the name (e.g., @{text haftmann@38441: "append_i_i_o"}). You can specify your own names as follows: haftmann@59377: \ haftmann@38441: haftmann@38811: code_pred %quote (modes: i \ i \ o \ bool as concat, haftmann@38811: o \ o \ i \ bool as split, haftmann@38811: i \ o \ i \ bool as suffix) append . haftmann@37613: haftmann@59377: subsection \Alternative introduction rules\ haftmann@38441: haftmann@59377: text \ haftmann@38441: Sometimes the introduction rules of an predicate are not executable haftmann@38441: because they contain non-executable constants or specific modes haftmann@38441: could not be inferred. It is also possible that the introduction haftmann@38441: rules yield a function that loops forever due to the execution in a haftmann@38441: depth-first search manner. Therefore, you can declare alternative haftmann@38441: introduction rules for predicates with the attribute @{attribute haftmann@38441: "code_pred_intro"}. For example, the transitive closure is defined haftmann@38441: by: haftmann@59377: \ haftmann@38441: haftmann@59377: text %quote \ haftmann@39682: @{lemma [source] "r a b \ tranclp r a b" by (fact tranclp.intros(1))}\\ haftmann@39682: @{lemma [source] "tranclp r a b \ r b c \ tranclp r a c" by (fact tranclp.intros(2))} haftmann@59377: \ haftmann@38441: haftmann@59377: text \ haftmann@38441: \noindent These rules do not suit well for executing the transitive haftmann@38441: closure with the mode @{text "(i \ o \ bool) \ i \ o \ bool"}, as haftmann@38441: the second rule will cause an infinite loop in the recursive call. haftmann@38441: This can be avoided using the following alternative rules which are haftmann@38441: declared to the predicate compiler by the attribute @{attribute haftmann@38441: "code_pred_intro"}: haftmann@59377: \ haftmann@38441: haftmann@37613: lemma %quote [code_pred_intro]: haftmann@39682: "r a b \ tranclp r a b" haftmann@39682: "r a b \ tranclp r b c \ tranclp r a c" haftmann@37613: by auto haftmann@38441: haftmann@59377: text \ haftmann@38441: \noindent After declaring all alternative rules for the transitive haftmann@38441: closure, you invoke @{command "code_pred"} as usual. As you have haftmann@38441: declared alternative rules for the predicate, you are urged to prove haftmann@38441: that these introduction rules are complete, i.e., that you can haftmann@38441: derive an elimination rule for the alternative rules: haftmann@59377: \ haftmann@38441: haftmann@37613: code_pred %quote tranclp haftmann@37613: proof - haftmann@37613: case tranclp haftmann@39682: from this converse_tranclpE [OF tranclp.prems] show thesis by metis haftmann@37613: qed haftmann@38441: haftmann@59377: text \ haftmann@38441: \noindent Alternative rules can also be used for constants that have haftmann@38441: not been defined inductively. For example, the lexicographic order haftmann@38441: which is defined as: haftmann@59377: \ haftmann@38441: haftmann@59377: text %quote \ haftmann@46515: @{thm [display] lexordp_def [of r]} haftmann@59377: \ haftmann@38441: haftmann@59377: text \ haftmann@38441: \noindent To make it executable, you can derive the following two haftmann@38441: rules and prove the elimination rule: haftmann@59377: \ haftmann@38441: haftmann@37613: lemma %quote [code_pred_intro]: haftmann@46515: "append xs (a # v) ys \ lexordp r xs ys" haftmann@46515: (*<*)unfolding lexordp_def by (auto simp add: append)(*>*) haftmann@37613: haftmann@37613: lemma %quote [code_pred_intro]: haftmann@46515: "append u (a # v) xs \ append u (b # w) ys \ r a b haftmann@46515: \ lexordp r xs ys" haftmann@46515: (*<*)unfolding lexordp_def append apply simp haftmann@38441: apply (rule disjI2) by auto(*>*) haftmann@37613: haftmann@46515: code_pred %quote lexordp haftmann@38441: (*<*)proof - haftmann@37613: fix r xs ys haftmann@46515: assume lexord: "lexordp r xs ys" haftmann@46515: assume 1: "\r' xs' ys' a v. r = r' \ xs = xs' \ ys = ys' haftmann@46515: \ append xs' (a # v) ys' \ thesis" haftmann@46515: assume 2: "\r' xs' ys' u a v b w. r = r' \ xs = xs' \ ys = ys' haftmann@46515: \ append u (a # v) xs' \ append u (b # w) ys' \ r' a b \ thesis" haftmann@37613: { haftmann@37613: assume "\a v. ys = xs @ a # v" haftmann@37613: from this 1 have thesis haftmann@46515: by (fastforce simp add: append) haftmann@37613: } moreover haftmann@37613: { haftmann@46515: assume "\u a b v w. r a b \ xs = u @ a # v \ ys = u @ b # w" haftmann@46515: from this 2 have thesis by (fastforce simp add: append) haftmann@37613: } moreover haftmann@37613: note lexord haftmann@37613: ultimately show thesis haftmann@46515: unfolding lexordp_def haftmann@46515: by fastforce haftmann@38441: qed(*>*) haftmann@38441: haftmann@38441: haftmann@59377: subsection \Options for values\ haftmann@38441: haftmann@59377: text \ haftmann@38441: In the presence of higher-order predicates, multiple modes for some haftmann@38441: predicate could be inferred that are not disambiguated by the haftmann@38441: pattern of the set comprehension. To disambiguate the modes for the haftmann@38441: arguments of a predicate, you can state the modes explicitly in the haftmann@38441: @{command "values"} command. Consider the simple predicate @{term haftmann@38441: "succ"}: haftmann@59377: \ haftmann@38441: haftmann@38441: inductive %quote succ :: "nat \ nat \ bool" where haftmann@37613: "succ 0 (Suc 0)" haftmann@37613: | "succ x y \ succ (Suc x) (Suc y)" haftmann@37613: haftmann@38441: code_pred %quote succ . haftmann@37613: haftmann@59377: text \ haftmann@38441: \noindent For this, the predicate compiler can infer modes @{text "o haftmann@38441: \ o \ bool"}, @{text "i \ o \ bool"}, @{text "o \ i \ bool"} and haftmann@38441: @{text "i \ i \ bool"}. The invocation of @{command "values"} haftmann@38441: @{text "{n. tranclp succ 10 n}"} loops, as multiple modes for the haftmann@38441: predicate @{text "succ"} are possible and here the first mode @{text haftmann@38441: "o \ o \ bool"} is chosen. To choose another mode for the argument, haftmann@38441: you can declare the mode for the argument between the @{command haftmann@38441: "values"} and the number of elements. haftmann@59377: \ haftmann@38441: wenzelm@61498: values %quote [mode: i \ o \ bool] 1 "{n. tranclp succ 10 n}" (*FIXME does not terminate for n\1*) haftmann@39065: values %quote [mode: o \ i \ bool] 1 "{n. tranclp succ n 10}" haftmann@37613: haftmann@38441: haftmann@59377: subsection \Embedding into functional code within Isabelle/HOL\ haftmann@38441: haftmann@59377: text \ haftmann@38441: To embed the computation of an inductive predicate into functions haftmann@38441: that are defined in Isabelle/HOL, you have a number of options: haftmann@38441: haftmann@38441: \begin{itemize} haftmann@38441: wenzelm@61498: \item You want to use the first-order predicate with the mode wenzelm@61498: where all arguments are input. Then you can use the predicate directly, e.g. haftmann@38441: wenzelm@61498: @{text [display] wenzelm@61498: \valid_suffix ys zs = wenzelm@61498: (if append [Suc 0, 2] ys zs then Some ys else None)\} haftmann@38441: wenzelm@61498: \item If you know that the execution returns only one value (it is wenzelm@61498: deterministic), then you can use the combinator @{term wenzelm@61498: "Predicate.the"}, e.g., a functional concatenation of lists is wenzelm@61498: defined with haftmann@38441: wenzelm@61498: @{term [display] "functional_concat xs ys = Predicate.the (append_i_i_o xs ys)"} haftmann@38441: wenzelm@61498: Note that if the evaluation does not return a unique value, it wenzelm@61498: raises a run-time error @{term "not_unique"}. haftmann@38441: haftmann@38441: \end{itemize} haftmann@59377: \ haftmann@38441: haftmann@38441: haftmann@59377: subsection \Further Examples\ haftmann@38441: haftmann@59377: text \ haftmann@38441: Further examples for compiling inductive predicates can be found in wenzelm@63680: \<^file>\~~/src/HOL/Predicate_Compile_Examples/Examples.thy\. There are haftmann@38441: also some examples in the Archive of Formal Proofs, notably in the haftmann@38441: @{text "POPLmark-deBruijn"} and the @{text "FeatherweightJava"} haftmann@38441: sessions. haftmann@59377: \ haftmann@38441: haftmann@37613: end haftmann@46515: