src/Doc/Codegen/Inductive_Predicate.thy
author nipkow
Fri, 11 Dec 2020 17:58:01 +0100
changeset 72884 50f18a822ee9
parent 72184 881bd98bddee
child 76987 4c275405faae
permissions -rw-r--r--
merged
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
69422
472af2d7835d clarified session dependencies: faster build_doc/build_release;
wenzelm
parents: 66453
diff changeset
     2
imports Setup
37613
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
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    19
section \<open>Inductive Predicates \label{sec:inductive}\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    20
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    21
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    22
  The \<open>predicate compiler\<close> 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
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 52753
diff changeset
    26
  @{cite "Berghofer-Bulwahn-Haftmann:2009:TPHOL"}.
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    27
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    28
  Consider the simple predicate \<^const>\<open>append\<close> given by these two
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    29
  introduction rules:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    30
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    31
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    32
text %quote \<open>
38441
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]}
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    35
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    36
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    37
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    38
  \noindent To invoke the compiler, simply use @{command_def "code_pred"}:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    39
\<close>
37613
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
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    43
text \<open>
38441
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
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    49
  output. Modes are similar to types, but use the notation \<open>i\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    50
  for input and \<open>o\<close> for output.
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    51
 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    52
  For \<^term>\<open>append\<close>, the compiler can infer the following modes:
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    53
  \begin{itemize}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    54
    \item \<open>i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    55
    \item \<open>i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    56
    \item \<open>o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool\<close>
38441
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"}:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    59
\<close>
38441
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
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    63
text \<open>\noindent outputs \<open>{[1, 2, 3, 4, 5]}\<close>, and\<close>
38441
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
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    67
text \<open>\noindent outputs \<open>{([], [2, 3]), ([2], [3]), ([2, 3], [])}\<close>.\<close>
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
    68
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    69
text \<open>
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:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    74
\<close>
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
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    79
text \<open>
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
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    84
  the suffix \<open>equation\<close>, and can be inspected with:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    85
\<close>
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
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    89
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    90
  \noindent More advanced options are described in the following subsections.
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    91
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    92
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    93
subsection \<open>Alternative names for functions\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    94
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    95
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    96
  By default, the functions generated from a predicate are named after
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
    97
  the predicate with the mode mangled into the name (e.g., \<open>append_i_i_o\<close>).  You can specify your own names as follows:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    98
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
    99
38811
c3511b112595 more xsymbols
haftmann
parents: 38508
diff changeset
   100
code_pred %quote (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool as concat,
c3511b112595 more xsymbols
haftmann
parents: 38508
diff changeset
   101
  o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as split,
c3511b112595 more xsymbols
haftmann
parents: 38508
diff changeset
   102
  i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as suffix) append .
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   103
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   104
subsection \<open>Alternative introduction rules\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   105
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   106
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   107
  Sometimes the introduction rules of an predicate are not executable
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   108
  because they contain non-executable constants or specific modes
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   109
  could not be inferred.  It is also possible that the introduction
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   110
  rules yield a function that loops forever due to the execution in a
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   111
  depth-first search manner.  Therefore, you can declare alternative
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   112
  introduction rules for predicates with the attribute @{attribute
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   113
  "code_pred_intro"}.  For example, the transitive closure is defined
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   114
  by:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   115
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   116
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   117
text %quote \<open>
39682
066e2d4d0de8 avoid fragile tranclp syntax; corrected resolution; corrected typo
haftmann
parents: 39065
diff changeset
   118
  @{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
   119
  @{lemma [source] "tranclp r a b \<Longrightarrow> r b c \<Longrightarrow> tranclp r a c" by (fact tranclp.intros(2))}
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   120
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   121
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   122
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   123
  \noindent These rules do not suit well for executing the transitive
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   124
  closure with the mode \<open>(i \<Rightarrow> o \<Rightarrow> bool) \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool\<close>, as
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   125
  the second rule will cause an infinite loop in the recursive call.
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   126
  This can be avoided using the following alternative rules which are
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   127
  declared to the predicate compiler by the attribute @{attribute
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   128
  "code_pred_intro"}:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   129
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   130
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   131
lemma %quote [code_pred_intro]:
39682
066e2d4d0de8 avoid fragile tranclp syntax; corrected resolution; corrected typo
haftmann
parents: 39065
diff changeset
   132
  "r a b \<Longrightarrow> tranclp r a b"
066e2d4d0de8 avoid fragile tranclp syntax; corrected resolution; corrected typo
haftmann
parents: 39065
diff changeset
   133
  "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
   134
by auto
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   135
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   136
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   137
  \noindent After declaring all alternative rules for the transitive
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   138
  closure, you invoke @{command "code_pred"} as usual.  As you have
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   139
  declared alternative rules for the predicate, you are urged to prove
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   140
  that these introduction rules are complete, i.e., that you can
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   141
  derive an elimination rule for the alternative rules:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   142
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   143
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   144
code_pred %quote tranclp
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   145
proof -
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   146
  case tranclp
39682
066e2d4d0de8 avoid fragile tranclp syntax; corrected resolution; corrected typo
haftmann
parents: 39065
diff changeset
   147
  from this converse_tranclpE [OF tranclp.prems] show thesis by metis
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   148
qed
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   149
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   150
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   151
  \noindent Alternative rules can also be used for constants that have
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   152
  not been defined inductively. For example, the lexicographic order
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   153
  which is defined as:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   154
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   155
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   156
text %quote \<open>
46515
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   157
  @{thm [display] lexordp_def [of r]}
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   158
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   159
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   160
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   161
  \noindent To make it executable, you can derive the following two
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   162
  rules and prove the elimination rule:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   163
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   164
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   165
lemma %quote [code_pred_intro]:
46515
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   166
  "append xs (a # v) ys \<Longrightarrow> lexordp r xs ys"
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   167
(*<*)unfolding lexordp_def by (auto simp add: append)(*>*)
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   168
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   169
lemma %quote [code_pred_intro]:
72184
881bd98bddee reversing all the lex crap
paulson <lp15@cam.ac.uk>
parents: 72172
diff changeset
   170
  "append u (a # v) xs \<Longrightarrow> append u (b # w) ys \<Longrightarrow> r a b
46515
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   171
  \<Longrightarrow> lexordp r xs ys"
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   172
(*<*)unfolding lexordp_def append apply simp
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   173
apply (rule disjI2) by auto(*>*)
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   174
46515
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   175
code_pred %quote lexordp
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   176
(*<*)proof -
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   177
  fix r xs ys
46515
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   178
  assume lexord: "lexordp r xs ys"
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   179
  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
   180
    \<Longrightarrow> append xs' (a # v) ys' \<Longrightarrow> thesis"
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   181
  assume 2: "\<And>r' xs' ys' u a v b w. r = r' \<Longrightarrow> xs = xs' \<Longrightarrow> ys = ys'
72184
881bd98bddee reversing all the lex crap
paulson <lp15@cam.ac.uk>
parents: 72172
diff changeset
   182
    \<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
   183
  {
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   184
    assume "\<exists>a v. ys = xs @ a # v"
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   185
    from this 1 have thesis
46515
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   186
        by (fastforce simp add: append)
37613
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
  {
72184
881bd98bddee reversing all the lex crap
paulson <lp15@cam.ac.uk>
parents: 72172
diff changeset
   189
    assume "\<exists>u a b v w. r a b \<and> xs = u @ a # v \<and> ys = u @ b # w"
46515
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   190
    from this 2 have thesis by (fastforce simp add: append)
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   191
  } moreover
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   192
  note lexord
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   193
  ultimately show thesis
46515
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   194
    unfolding lexordp_def
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   195
    by fastforce
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   196
qed(*>*)
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   197
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   198
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   199
subsection \<open>Options for values\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   200
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   201
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   202
  In the presence of higher-order predicates, multiple modes for some
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   203
  predicate could be inferred that are not disambiguated by the
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   204
  pattern of the set comprehension.  To disambiguate the modes for the
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   205
  arguments of a predicate, you can state the modes explicitly in the
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   206
  @{command "values"} command.  Consider the simple predicate \<^term>\<open>succ\<close>:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   207
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   208
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   209
inductive %quote succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   210
  "succ 0 (Suc 0)"
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   211
| "succ x y \<Longrightarrow> succ (Suc x) (Suc y)"
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   212
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   213
code_pred %quote succ .
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   214
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   215
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   216
  \noindent For this, the predicate compiler can infer modes \<open>o
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   217
  \<Rightarrow> o \<Rightarrow> bool\<close>, \<open>i \<Rightarrow> o \<Rightarrow> bool\<close>, \<open>o \<Rightarrow> i \<Rightarrow> bool\<close> and
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   218
  \<open>i \<Rightarrow> i \<Rightarrow> bool\<close>.  The invocation of @{command "values"}
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   219
  \<open>{n. tranclp succ 10 n}\<close> loops, as multiple modes for the
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   220
  predicate \<open>succ\<close> are possible and here the first mode \<open>o \<Rightarrow> o \<Rightarrow> bool\<close> is chosen. To choose another mode for the argument,
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   221
  you can declare the mode for the argument between the @{command
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   222
  "values"} and the number of elements.
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   223
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   224
61498
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   225
values %quote [mode: i \<Rightarrow> o \<Rightarrow> bool] 1 "{n. tranclp succ 10 n}" (*FIXME does not terminate for n\<ge>1*)
39065
16a2ed09217a set depth to 1
haftmann
parents: 38811
diff changeset
   226
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
   227
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   228
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   229
subsection \<open>Embedding into functional code within Isabelle/HOL\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   230
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   231
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   232
  To embed the computation of an inductive predicate into functions
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   233
  that are defined in Isabelle/HOL, you have a number of options:
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   234
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   235
  \begin{itemize}
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   236
61498
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   237
  \item You want to use the first-order predicate with the mode
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   238
    where all arguments are input. Then you can use the predicate directly, e.g.
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   239
61498
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   240
    @{text [display]
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   241
\<open>valid_suffix ys zs =
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   242
  (if append [Suc 0, 2] ys zs then Some ys else None)\<close>}
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   243
61498
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   244
  \item If you know that the execution returns only one value (it is
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   245
    deterministic), then you can use the combinator \<^term>\<open>Predicate.the\<close>, e.g., a functional concatenation of lists is
61498
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   246
    defined with
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   247
61498
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   248
    @{term [display] "functional_concat xs ys = Predicate.the (append_i_i_o xs ys)"}
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   249
61498
32a20d7b3ee4 tuned document;
wenzelm
parents: 59377
diff changeset
   250
    Note that if the evaluation does not return a unique value, it
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   251
    raises a run-time error \<^term>\<open>not_unique\<close>.
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   252
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   253
  \end{itemize}
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   254
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   255
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   256
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   257
subsection \<open>Further Examples\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   258
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   259
text \<open>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   260
  Further examples for compiling inductive predicates can be found in
63680
6e1e8b5abbfa more symbols;
wenzelm
parents: 61498
diff changeset
   261
  \<^file>\<open>~~/src/HOL/Predicate_Compile_Examples/Examples.thy\<close>.  There are
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   262
  also some examples in the Archive of Formal Proofs, notably in the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   263
  \<open>POPLmark-deBruijn\<close> and the \<open>FeatherweightJava\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   264
  sessions.
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   265
\<close>
38441
2fffd5ac487f tuned section on predicate compiler
haftmann
parents: 38437
diff changeset
   266
37613
355ec1b521e6 split off predicate compiler into separate theory
haftmann
parents:
diff changeset
   267
end
46515
2a0e1bcf713c adjusted to set type constructor
haftmann
parents: 39682
diff changeset
   268