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