CookBook/Package/Ind_Examples.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 17 Dec 2008 05:08:33 +0000
changeset 59 b5914f3c643c
parent 38 e21b2f888fa2
child 88 ebbd0dd008c8
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     1
theory Ind_Examples
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     2
imports Main
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     3
begin
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     4
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     5
section{* Examples of inductive definitions *}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     6
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     7
text {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     8
\label{sec:ind-examples}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     9
In this section, we will give three examples showing how to define inductive
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    10
predicates by hand and prove characteristic properties such as introduction
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    11
rules and an induction rule. From these examples, we will then figure out a
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    12
general method for defining inductive predicates, which will be described in
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    13
\S\ref{sec:ind-general-method}. This description will serve as a basis for the
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    14
actual implementation to be developed in \S\ref{sec:ind-interface} -- \S\ref{sec:ind-proofs}.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    15
It should be noted that our aim in this section is not to write proofs that
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    16
are as beautiful as possible, but as close as possible to the ML code producing
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    17
the proofs that we will develop later.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    18
As a first example, we consider the \emph{transitive closure} @{text "trcl R"}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    19
of a relation @{text R}. It is characterized by the following
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    20
two introduction rules
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    21
\[
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    22
\begin{array}{l}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    23
@{term "trcl R x x"} \\
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    24
@{term [mode=no_brackets] "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    25
\end{array}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    26
\]
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    27
Note that the @{text trcl} predicate has two different kinds of parameters: the
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    28
first parameter @{text R} stays \emph{fixed} throughout the definition, whereas
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    29
the second and third parameter changes in the ``recursive call''.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    30
Since an inductively defined predicate is the \emph{least} predicate closed under
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    31
a collection of introduction rules, we define the predicate @{text "trcl R x y"} in
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    32
such a way that it holds if and only if @{text "P x y"} holds for every predicate
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    33
@{text P} closed under the above rules. This gives rise to a definition containing
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    34
a universal quantifier over the predicate @{text P}:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    35
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    36
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    37
definition "trcl \<equiv> \<lambda>R x y.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    38
  \<forall>P. (\<forall>x. P x x) \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P x y"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    39
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    40
text {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    41
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    42
Since the predicate @{term "trcl R x y"} yields an element of the type of object
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    43
level truth values @{text bool}, the meta-level implications @{text "\<Longrightarrow>"} in the
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    44
above introduction rules have to be converted to object-level implications
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    45
@{text "\<longrightarrow>"}. Moreover, we use object-level universal quantifiers @{text "\<forall>"}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    46
rather than meta-level universal quantifiers @{text "\<And>"} for quantifying over
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    47
the variable parameters of the introduction rules. Isabelle already offers some
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    48
infrastructure for converting between meta-level and object-level connectives,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    49
which we will use later on.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    50
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    51
With this definition of the transitive closure, the proof of the (weak) induction
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    52
theorem is almost immediate. It suffices to convert all the meta-level connectives
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    53
in the induction rule to object-level connectives using the @{text atomize} proof
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    54
method, expand the definition of @{text trcl}, eliminate the universal quantifier
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    55
contained in it, and then solve the goal by assumption.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    56
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    57
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    58
lemma trcl_induct:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    59
  assumes trcl: "trcl R x y"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    60
  shows "(\<And>x. P x x) \<Longrightarrow> (\<And>x y z. R x y \<Longrightarrow> P y z \<Longrightarrow> P x z) \<Longrightarrow> P x y"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    61
  apply (atomize (full))
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    62
  apply (cut_tac trcl)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    63
  apply (unfold trcl_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    64
  apply (drule spec [where x=P])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    65
  apply assumption
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    66
  done
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    67
(*<*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    68
lemma "trcl R x x"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    69
  apply (unfold trcl_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    70
  apply (rule allI impI)+
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    71
(*>*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    72
txt {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    73
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    74
The above induction rule is \emph{weak} in the sense that the induction step may
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    75
only be proved using the assumptions @{term "R x y"} and @{term "P y z"}, but not
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    76
using the additional assumption \mbox{@{term "trcl R y z"}}. A stronger induction rule
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    77
containing this additional assumption can be derived from the weaker one with the
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    78
help of the introduction rules for @{text trcl}.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    79
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    80
We now turn to the proofs of the introduction rules, which are slightly more complicated.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    81
In order to prove the first introduction rule, we again unfold the definition and
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    82
then apply the introdution rules for @{text "\<forall>"} and @{text "\<longrightarrow>"} as often as possible.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    83
We then end up in a proof state of the following form:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    84
@{subgoals [display]}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    85
The two assumptions correspond to the introduction rules, where @{term "trcl R"} has been
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    86
replaced by @{term "P"}. Thus, all we have to do is to eliminate the universal quantifier
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    87
in front of the first assumption, and then solve the goal by assumption:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    88
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    89
(*<*)oops(*>*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    90
lemma trcl_base: "trcl R x x"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    91
  apply (unfold trcl_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    92
  apply (rule allI impI)+
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    93
  apply (drule spec)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    94
  apply assumption
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    95
  done
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    96
(*<*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    97
lemma "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    98
  apply (unfold trcl_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    99
  apply (rule allI impI)+
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   100
(*>*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   101
txt {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   102
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   103
Since the second introduction rule has premises, its proof is not as easy as the previous
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   104
one. After unfolding the definitions and applying the introduction rules for @{text "\<forall>"}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   105
and @{text "\<longrightarrow>"}, we get the proof state
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   106
@{subgoals [display]}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   107
The third and fourth assumption corresponds to the first and second introduction rule,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   108
respectively, whereas the first and second assumption corresponds to the premises of
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   109
the introduction rule. Since we want to prove the second introduction rule, we apply
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   110
the fourth assumption to the goal @{term "P x z"}. In order for the assumption to
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   111
be applicable, we have to eliminate the universal quantifiers and turn the object-level
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   112
implications into meta-level ones. This can be accomplished using the @{text rule_format}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   113
attribute. Applying the assumption produces two new subgoals, which can be solved using
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   114
the first and second assumption. The second assumption again involves a quantifier and
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   115
implications that have to be eliminated before it can be applied. To avoid problems
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   116
with higher order unification, it is advisable to provide an instantiation for the
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   117
universally quantified predicate variable in the assumption.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   118
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   119
(*<*)oops(*>*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   120
lemma trcl_step: "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   121
  apply (unfold trcl_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   122
  apply (rule allI impI)+
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   123
  proof -
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   124
    case goal1
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   125
    show ?case
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   126
      apply (rule goal1(4) [rule_format])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   127
      apply (rule goal1(1))
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   128
      apply (rule goal1(2) [THEN spec [where x=P], THEN mp, THEN mp,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   129
	OF goal1(3-4)])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   130
      done
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   131
  qed
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   132
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   133
text {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   134
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   135
This method of defining inductive predicates easily generalizes to mutually inductive
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   136
predicates, like the predicates @{text even} and @{text odd} characterized by the
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   137
following introduction rules:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   138
\[
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   139
\begin{array}{l}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   140
@{term "even (0::nat)"} \\
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   141
@{term "odd m \<Longrightarrow> even (Suc m)"} \\
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   142
@{term "even m \<Longrightarrow> odd (Suc m)"}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   143
\end{array}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   144
\]
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   145
Since the predicates are mutually inductive, each of the definitions contain two
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   146
quantifiers over the predicates @{text P} and @{text Q}.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   147
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   148
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   149
definition "even \<equiv> \<lambda>n.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   150
  \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m)) \<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> P n"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   151
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   152
definition "odd \<equiv> \<lambda>n.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   153
  \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m)) \<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> Q n"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   154
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   155
text {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   156
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   157
For proving the induction rule, we use exactly the same technique as in the transitive
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   158
closure example:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   159
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   160
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   161
lemma even_induct:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   162
  assumes even: "even n"
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 32
diff changeset
   163
  shows "P 0 \<Longrightarrow> 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 32
diff changeset
   164
             (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   165
  apply (atomize (full))
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   166
  apply (cut_tac even)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   167
  apply (unfold even_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   168
  apply (drule spec [where x=P])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   169
  apply (drule spec [where x=Q])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   170
  apply assumption
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   171
  done
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   172
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   173
text {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   174
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   175
A similar induction rule having @{term "Q n"} as a conclusion can be proved for
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   176
the @{text odd} predicate.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   177
The proofs of the introduction rules are also very similar to the ones in the
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   178
previous example. We only show the proof of the second introduction rule,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   179
since it is almost the same as the one for the third introduction rule,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   180
and the proof of the first rule is trivial.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   181
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   182
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   183
lemma evenS: "odd m \<Longrightarrow> even (Suc m)"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   184
  apply (unfold odd_def even_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   185
  apply (rule allI impI)+
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   186
  proof -
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   187
    case goal1
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   188
    show ?case
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   189
      apply (rule goal1(3) [rule_format])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   190
      apply (rule goal1(1) [THEN spec [where x=P], THEN spec [where x=Q],
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   191
	THEN mp, THEN mp, THEN mp, OF goal1(2-4)])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   192
      done
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   193
  qed
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   194
(*<*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   195
lemma even0: "even 0"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   196
  apply (unfold even_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   197
  apply (rule allI impI)+
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   198
  apply assumption
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   199
  done
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   200
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   201
lemma oddS: "even m \<Longrightarrow> odd (Suc m)"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   202
  apply (unfold odd_def even_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   203
  apply (rule allI impI)+
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   204
  proof -
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   205
    case goal1
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   206
    show ?case
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   207
      apply (rule goal1(4) [rule_format])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   208
      apply (rule goal1(1) [THEN spec [where x=P], THEN spec [where x=Q],
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   209
	THEN mp, THEN mp, THEN mp, OF goal1(2-4)])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   210
      done
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   211
  qed
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   212
(*>*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   213
text {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   214
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   215
As a final example, we will consider the definition of the accessible
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   216
part of a relation @{text R} characterized by the introduction rule
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   217
\[
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   218
@{term "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   219
\]
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   220
whose premise involves a universal quantifier and an implication. The
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   221
definition of @{text accpart} is as follows:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   222
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   223
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   224
definition "accpart \<equiv> \<lambda>R x. \<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P x"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   225
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   226
text {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   227
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   228
The proof of the induction theorem is again straightforward:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   229
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   230
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   231
lemma accpart_induct:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   232
  assumes acc: "accpart R x"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   233
  shows "(\<And>x. (\<And>y. R y x \<Longrightarrow> P y) \<Longrightarrow> P x) \<Longrightarrow> P x"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   234
  apply (atomize (full))
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   235
  apply (cut_tac acc)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   236
  apply (unfold accpart_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   237
  apply (drule spec [where x=P])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   238
  apply assumption
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   239
  done
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   240
(*<*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   241
lemma accpartI: "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   242
  apply (unfold accpart_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   243
  apply (rule allI impI)+
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   244
(*>*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   245
txt {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   246
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   247
Proving the introduction rule is a little more complicated, due to the quantifier
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   248
and the implication in the premise. We first convert the meta-level universal quantifier
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   249
and implication to their object-level counterparts. Unfolding the definition of
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   250
@{text accpart} and applying the introduction rules for @{text "\<forall>"} and @{text "\<longrightarrow>"}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   251
yields the following proof state:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   252
@{subgoals [display]}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   253
Applying the second assumption produces a proof state with the new local assumption
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   254
@{term "R y x"}, which will then be used to solve the goal @{term "P y"} using the
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   255
first assumption.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   256
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   257
(*<*)oops(*>*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   258
lemma accpartI: "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   259
  apply (unfold accpart_def)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   260
  apply (rule allI impI)+
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   261
  proof -
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   262
    case goal1
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   263
    note goal1' = this
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   264
    show ?case
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   265
      apply (rule goal1'(2) [rule_format])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   266
      proof -
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   267
        case goal1
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   268
        show ?case
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   269
	  apply (rule goal1'(1) [OF goal1, THEN spec [where x=P],
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   270
            THEN mp, OF goal1'(2)])
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   271
	  done
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   272
      qed
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   273
    qed
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   274
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   275
end