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