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