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