ProgTutorial/Package/Ind_Prelims.thy
changeset 565 cecd7a941885
parent 562 daf404920ab9
child 573 321e220a6baa
equal deleted inserted replaced
564:6e2479089226 565:cecd7a941885
     1 theory Ind_Prelims
     1 theory Ind_Prelims
     2 imports Ind_Intro 
     2 imports Ind_Intro 
     3 begin
     3 begin
     4 
     4 
     5 section{* Preliminaries *}
     5 section\<open>Preliminaries\<close>
     6   
     6   
     7 text {*
     7 text \<open>
     8   The user will just give a specification of inductive predicate(s) and
     8   The user will just give a specification of inductive predicate(s) and
     9   expects from the package to produce a convenient reasoning
     9   expects from the package to produce a convenient reasoning
    10   infrastructure. This infrastructure needs to be derived from the definition
    10   infrastructure. This infrastructure needs to be derived from the definition
    11   that correspond to the specified predicate(s). Before we start with
    11   that correspond to the specified predicate(s). Before we start with
    12   explaining all parts of the package, let us first give some examples showing
    12   explaining all parts of the package, let us first give some examples showing
    16   writing a package for Isabelle. The aim in this section is \emph{not} to
    16   writing a package for Isabelle. The aim in this section is \emph{not} to
    17   write proofs that are as beautiful as possible, but as close as possible to
    17   write proofs that are as beautiful as possible, but as close as possible to
    18   the ML-implementation we will develop in later sections.
    18   the ML-implementation we will develop in later sections.
    19 
    19 
    20 
    20 
    21   We first consider the transitive closure of a relation @{text R}. The 
    21   We first consider the transitive closure of a relation \<open>R\<close>. The 
    22   ``pencil-and-paper'' specification for the transitive closure is:
    22   ``pencil-and-paper'' specification for the transitive closure is:
    23 
    23 
    24   \begin{center}\small
    24   \begin{center}\small
    25   @{prop[mode=Axiom] "trcl R x x"} \hspace{5mm}
    25   @{prop[mode=Axiom] "trcl R x x"} \hspace{5mm}
    26   @{prop[mode=Rule] "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"}
    26   @{prop[mode=Rule] "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"}
    27   \end{center}
    27   \end{center}
    28 
    28 
    29   As mentioned before, the package has to make an appropriate definition for
    29   As mentioned before, the package has to make an appropriate definition for
    30   @{term "trcl"}. Since an inductively defined predicate is the least
    30   @{term "trcl"}. Since an inductively defined predicate is the least
    31   predicate closed under a collection of introduction rules, the predicate
    31   predicate closed under a collection of introduction rules, the predicate
    32   @{text "trcl R x y"} can be defined so that it holds if and only if @{text
    32   \<open>trcl R x y\<close> can be defined so that it holds if and only if \<open>P x y\<close> holds for every predicate \<open>P\<close> closed under the rules
    33   "P x y"} holds for every predicate @{text P} closed under the rules
       
    34   above. This gives rise to the definition
    33   above. This gives rise to the definition
    35 *}
    34 \<close>
    36 
    35 
    37 definition "trcl \<equiv> 
    36 definition "trcl \<equiv> 
    38      \<lambda>R x y. \<forall>P. (\<forall>x. P x x) 
    37      \<lambda>R x y. \<forall>P. (\<forall>x. P x x) 
    39                   \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P x y"
    38                   \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P x y"
    40 
    39 
    41 text {*
    40 text \<open>
    42   Note we have to use the object implication @{text "\<longrightarrow>"} and object
    41   Note we have to use the object implication \<open>\<longrightarrow>\<close> and object
    43   quantification @{text "\<forall>"} for stating this definition (there is no other
    42   quantification \<open>\<forall>\<close> for stating this definition (there is no other
    44   way for definitions in HOL). However, the introduction rules and induction
    43   way for definitions in HOL). However, the introduction rules and induction
    45   principles associated with the transitive closure should use the
    44   principles associated with the transitive closure should use the
    46   meta-connectives, since they simplify the reasoning for the user.
    45   meta-connectives, since they simplify the reasoning for the user.
    47 
    46 
    48 
    47 
    49   With this definition, the proof of the induction principle for @{term trcl}
    48   With this definition, the proof of the induction principle for @{term trcl}
    50   is almost immediate. It suffices to convert all the meta-level
    49   is almost immediate. It suffices to convert all the meta-level
    51   connectives in the lemma to object-level connectives using the
    50   connectives in the lemma to object-level connectives using the
    52   proof method @{text atomize} (Line 4 below), expand the definition of @{term trcl}
    51   proof method \<open>atomize\<close> (Line 4 below), expand the definition of @{term trcl}
    53   (Line 5 and 6), eliminate the universal quantifier contained in it (Line~7),
    52   (Line 5 and 6), eliminate the universal quantifier contained in it (Line~7),
    54   and then solve the goal by @{text assumption} (Line 8).
    53   and then solve the goal by \<open>assumption\<close> (Line 8).
    55 
    54 
    56 *}
    55 \<close>
    57 
    56 
    58 lemma %linenos trcl_induct:
    57 lemma %linenos trcl_induct:
    59 assumes "trcl R x y"
    58 assumes "trcl R x y"
    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"
    59 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"
    61 apply(atomize (full))
    60 apply(atomize (full))
    63 apply(unfold trcl_def)
    62 apply(unfold trcl_def)
    64 apply(drule spec[where x=P])
    63 apply(drule spec[where x=P])
    65 apply(assumption)
    64 apply(assumption)
    66 done
    65 done
    67 
    66 
    68 text {*
    67 text \<open>
    69   The proofs for the introduction rules are slightly more complicated. 
    68   The proofs for the introduction rules are slightly more complicated. 
    70   For the first one, we need to prove the following lemma:
    69   For the first one, we need to prove the following lemma:
    71 *}
    70 \<close>
    72 
    71 
    73 lemma %linenos trcl_base: 
    72 lemma %linenos trcl_base: 
    74 shows "trcl R x x"
    73 shows "trcl R x x"
    75 apply(unfold trcl_def)
    74 apply(unfold trcl_def)
    76 apply(rule allI impI)+
    75 apply(rule allI impI)+
    77 apply(drule spec)
    76 apply(drule spec)
    78 apply(assumption)
    77 apply(assumption)
    79 done
    78 done
    80 
    79 
    81 text {*
    80 text \<open>
    82   We again unfold first the definition and apply introduction rules 
    81   We again unfold first the definition and apply introduction rules 
    83   for @{text "\<forall>"} and @{text "\<longrightarrow>"} as often as possible (Lines 3 and 4).
    82   for \<open>\<forall>\<close> and \<open>\<longrightarrow>\<close> as often as possible (Lines 3 and 4).
    84   We then end up in the goal state:
    83   We then end up in the goal state:
    85 *}
    84 \<close>
    86 
    85 
    87 (*<*)lemma "trcl R x x"
    86 (*<*)lemma "trcl R x x"
    88 apply (unfold trcl_def)
    87 apply (unfold trcl_def)
    89 apply (rule allI impI)+(*>*)
    88 apply (rule allI impI)+(*>*)
    90 txt {* @{subgoals [display]} *}
    89 txt \<open>@{subgoals [display]}\<close>
    91 (*<*)oops(*>*)
    90 (*<*)oops(*>*)
    92 
    91 
    93 text {*
    92 text \<open>
    94   The two assumptions come from the definition of @{term trcl} and correspond
    93   The two assumptions come from the definition of @{term trcl} and correspond
    95   to the introduction rules. Thus, all we have to do is to eliminate the
    94   to the introduction rules. Thus, all we have to do is to eliminate the
    96   universal quantifier in front of the first assumption (Line 5), and then
    95   universal quantifier in front of the first assumption (Line 5), and then
    97   solve the goal by @{text assumption} (Line 6).
    96   solve the goal by \<open>assumption\<close> (Line 6).
    98 *}
    97 \<close>
    99 
    98 
   100 text {*
    99 text \<open>
   101   Next we have to show that the second introduction rule also follows from the
   100   Next we have to show that the second introduction rule also follows from the
   102   definition.  Since this rule has premises, the proof is a bit more
   101   definition.  Since this rule has premises, the proof is a bit more
   103   involved. After unfolding the definitions and applying the introduction
   102   involved. After unfolding the definitions and applying the introduction
   104   rules for @{text "\<forall>"} and @{text "\<longrightarrow>"}
   103   rules for \<open>\<forall>\<close> and \<open>\<longrightarrow>\<close>
   105 *}
   104 \<close>
   106 
   105 
   107 lemma trcl_step: 
   106 lemma trcl_step: 
   108 shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
   107 shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
   109 apply (unfold trcl_def)
   108 apply (unfold trcl_def)
   110 apply (rule allI impI)+
   109 apply (rule allI impI)+
   111 
   110 
   112 txt {* 
   111 txt \<open>
   113   we obtain the goal state
   112   we obtain the goal state
   114 
   113 
   115   @{subgoals [display]} 
   114   @{subgoals [display]} 
   116 
   115 
   117   To see better where we are, let us explicitly name the assumptions 
   116   To see better where we are, let us explicitly name the assumptions 
   118   by starting a subproof.
   117   by starting a subproof.
   119 *}
   118 \<close>
   120 
   119 
   121 proof -
   120 proof -
   122   case (goal1 P)
   121   case (goal1 P)
   123   have p1: "R x y" by fact
   122   have p1: "R x y" by fact
   124   have p2: "\<forall>P. (\<forall>x. P x x) 
   123   have p2: "\<forall>P. (\<forall>x. P x x) 
   125                   \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P y z" by fact
   124                   \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P y z" by fact
   126   have r1: "\<forall>x. P x x" by fact
   125   have r1: "\<forall>x. P x x" by fact
   127   have r2: "\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z" by fact
   126   have r2: "\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z" by fact
   128   show "P x z"
   127   show "P x z"
   129   
   128   
   130 txt {*
   129 txt \<open>
   131   The assumptions @{text "p1"} and @{text "p2"} correspond to the premises of
   130   The assumptions \<open>p1\<close> and \<open>p2\<close> correspond to the premises of
   132   the second introduction rule (unfolded); the assumptions @{text "r1"} and @{text "r2"}
   131   the second introduction rule (unfolded); the assumptions \<open>r1\<close> and \<open>r2\<close>
   133   come from the definition of @{term trcl}. We apply @{text "r2"} to the goal
   132   come from the definition of @{term trcl}. We apply \<open>r2\<close> to the goal
   134   @{term "P x z"}. In order for this assumption to be applicable as a rule, we
   133   @{term "P x z"}. In order for this assumption to be applicable as a rule, we
   135   have to eliminate the universal quantifier and turn the object-level
   134   have to eliminate the universal quantifier and turn the object-level
   136   implications into meta-level ones. This can be accomplished using the @{text
   135   implications into meta-level ones. This can be accomplished using the \<open>rule_format\<close> attribute. So we continue the proof with:
   137   rule_format} attribute. So we continue the proof with:
   136 
   138 
   137 \<close>
   139 *}
       
   140 
   138 
   141     apply (rule r2[rule_format])
   139     apply (rule r2[rule_format])
   142 
   140 
   143 txt {*
   141 txt \<open>
   144   This gives us two new subgoals
   142   This gives us two new subgoals
   145 
   143 
   146   @{subgoals [display]} 
   144   @{subgoals [display]} 
   147 
   145 
   148   which can be solved using assumptions @{text p1} and @{text p2}. The latter
   146   which can be solved using assumptions \<open>p1\<close> and \<open>p2\<close>. The latter
   149   involves a quantifier and implications that have to be eliminated before it
   147   involves a quantifier and implications that have to be eliminated before it
   150   can be applied. To avoid potential problems with higher-order unification,
   148   can be applied. To avoid potential problems with higher-order unification,
   151   we explicitly instantiate the quantifier to @{text "P"} and also match
   149   we explicitly instantiate the quantifier to \<open>P\<close> and also match
   152   explicitly the implications with @{text "r1"} and @{text "r2"}. This gives
   150   explicitly the implications with \<open>r1\<close> and \<open>r2\<close>. This gives
   153   the proof:
   151   the proof:
   154 *}
   152 \<close>
   155 
   153 
   156     apply(rule p1)
   154     apply(rule p1)
   157     apply(rule p2[THEN spec[where x=P], THEN mp, THEN mp, OF r1, OF r2])
   155     apply(rule p2[THEN spec[where x=P], THEN mp, THEN mp, OF r1, OF r2])
   158     done
   156     done
   159 qed
   157 qed
   160 
   158 
   161 text {*
   159 text \<open>
   162   Now we are done. It might be surprising that we are not using the automatic
   160   Now we are done. It might be surprising that we are not using the automatic
   163   tactics available in Isabelle/HOL for proving this lemmas. After all @{text
   161   tactics available in Isabelle/HOL for proving this lemmas. After all \<open>blast\<close> would easily dispense of it.
   164   "blast"} would easily dispense of it.
   162 \<close>
   165 *}
       
   166 
   163 
   167 lemma trcl_step_blast: 
   164 lemma trcl_step_blast: 
   168 shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
   165 shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
   169 apply(unfold trcl_def)
   166 apply(unfold trcl_def)
   170 apply(blast)
   167 apply(blast)
   171 done
   168 done
   172 term "even"
   169 term "even"
   173 text {*
   170 text \<open>
   174   Experience has shown that it is generally a bad idea to rely heavily on
   171   Experience has shown that it is generally a bad idea to rely heavily on
   175   @{text blast}, @{text auto} and the like in automated proofs. The reason is
   172   \<open>blast\<close>, \<open>auto\<close> and the like in automated proofs. The reason is
   176   that you do not have precise control over them (the user can, for example,
   173   that you do not have precise control over them (the user can, for example,
   177   declare new intro- or simplification rules that can throw automatic tactics
   174   declare new intro- or simplification rules that can throw automatic tactics
   178   off course) and also it is very hard to debug proofs involving automatic
   175   off course) and also it is very hard to debug proofs involving automatic
   179   tactics whenever something goes wrong. Therefore if possible, automatic 
   176   tactics whenever something goes wrong. Therefore if possible, automatic 
   180   tactics in packages should be avoided or be constrained sufficiently.
   177   tactics in packages should be avoided or be constrained sufficiently.
   181 
   178 
   182   The method of defining inductive predicates by impredicative quantification
   179   The method of defining inductive predicates by impredicative quantification
   183   also generalises to mutually inductive predicates. The next example defines
   180   also generalises to mutually inductive predicates. The next example defines
   184   the predicates @{text even} and @{text odd} given by
   181   the predicates \<open>even\<close> and \<open>odd\<close> given by
   185 
   182 
   186   \begin{center}\small
   183   \begin{center}\small
   187   @{prop[mode=Axiom] "even (0::nat)"} \hspace{5mm}
   184   @{prop[mode=Axiom] "even (0::nat)"} \hspace{5mm}
   188   @{prop[mode=Rule] "odd n \<Longrightarrow> even (Suc n)"} \hspace{5mm}
   185   @{prop[mode=Rule] "odd n \<Longrightarrow> even (Suc n)"} \hspace{5mm}
   189   @{prop[mode=Rule] "even n \<Longrightarrow> odd (Suc n)"}
   186   @{prop[mode=Rule] "even n \<Longrightarrow> odd (Suc n)"}
   190   \end{center}
   187   \end{center}
   191 
   188 
   192   Since the predicates @{term even} and @{term odd} are mutually inductive, each 
   189   Since the predicates @{term even} and @{term odd} are mutually inductive, each 
   193   corresponding definition must quantify over both predicates (we name them 
   190   corresponding definition must quantify over both predicates (we name them 
   194   below @{text "P"} and @{text "Q"}).
   191   below \<open>P\<close> and \<open>Q\<close>).
   195 *}
   192 \<close>
   196 
   193 
   197 hide_const even
   194 hide_const even
   198 hide_const odd
   195 hide_const odd
   199 
   196 
   200 definition "even \<equiv>
   197 definition "even \<equiv>
   203 
   200 
   204 definition "odd \<equiv>
   201 definition "odd \<equiv>
   205   \<lambda>n. \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m)) 
   202   \<lambda>n. \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m)) 
   206                  \<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> Q n"
   203                  \<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> Q n"
   207 
   204 
   208 text {*
   205 text \<open>
   209   For proving the induction principles, we use exactly the same technique 
   206   For proving the induction principles, we use exactly the same technique 
   210   as in the transitive closure example, namely:
   207   as in the transitive closure example, namely:
   211 *}
   208 \<close>
   212 
   209 
   213 lemma even_induct:
   210 lemma even_induct:
   214 assumes "even n"
   211 assumes "even n"
   215 shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
   212 shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
   216 apply(atomize (full))
   213 apply(atomize (full))
   219 apply(drule spec[where x=P])
   216 apply(drule spec[where x=P])
   220 apply(drule spec[where x=Q])
   217 apply(drule spec[where x=Q])
   221 apply(assumption)
   218 apply(assumption)
   222 done
   219 done
   223 
   220 
   224 text {*
   221 text \<open>
   225   The only difference with the proof @{text "trcl_induct"} is that we have to
   222   The only difference with the proof \<open>trcl_induct\<close> is that we have to
   226   instantiate here two universal quantifiers.  We omit the other induction
   223   instantiate here two universal quantifiers.  We omit the other induction
   227   principle that has @{prop "even n"} as premise and @{term "Q n"} as conclusion.  
   224   principle that has @{prop "even n"} as premise and @{term "Q n"} as conclusion.  
   228   The proofs of the introduction rules are also very similar to the ones in 
   225   The proofs of the introduction rules are also very similar to the ones in 
   229   the @{text "trcl"}-example. We only show the proof of the second introduction 
   226   the \<open>trcl\<close>-example. We only show the proof of the second introduction 
   230   rule.
   227   rule.
   231 *}
   228 \<close>
   232 
   229 
   233 lemma %linenos evenS: 
   230 lemma %linenos evenS: 
   234 shows "odd m \<Longrightarrow> even (Suc m)"
   231 shows "odd m \<Longrightarrow> even (Suc m)"
   235 apply (unfold odd_def even_def)
   232 apply (unfold odd_def even_def)
   236 apply (rule allI impI)+
   233 apply (rule allI impI)+
   246     apply(rule p1[THEN spec[where x=P], THEN spec[where x=Q],
   243     apply(rule p1[THEN spec[where x=P], THEN spec[where x=Q],
   247 	           THEN mp, THEN mp, THEN mp, OF r1, OF r2, OF r3])
   244 	           THEN mp, THEN mp, THEN mp, OF r1, OF r2, OF r3])
   248     done
   245     done
   249 qed
   246 qed
   250 
   247 
   251 text {*
   248 text \<open>
   252   The interesting lines are 7 to 15. Again, the assumptions fall into two categories:
   249   The interesting lines are 7 to 15. Again, the assumptions fall into two categories:
   253   @{text p1} corresponds to the premise of the introduction rule; @{text "r1"}
   250   \<open>p1\<close> corresponds to the premise of the introduction rule; \<open>r1\<close>
   254   to @{text "r3"} come from the definition of @{text "even"}.
   251   to \<open>r3\<close> come from the definition of \<open>even\<close>.
   255   In Line 13, we apply the assumption @{text "r2"} (since we prove the second
   252   In Line 13, we apply the assumption \<open>r2\<close> (since we prove the second
   256   introduction rule). In Lines 14 and 15 we apply assumption @{text "p1"} (if
   253   introduction rule). In Lines 14 and 15 we apply assumption \<open>p1\<close> (if
   257   the second introduction rule had more premises we have to do that for all
   254   the second introduction rule had more premises we have to do that for all
   258   of them). In order for this assumption to be applicable, the quantifiers
   255   of them). In order for this assumption to be applicable, the quantifiers
   259   need to be instantiated and then also the implications need to be resolved
   256   need to be instantiated and then also the implications need to be resolved
   260   with the other rules.
   257   with the other rules.
   261 
   258 
   262   Next we define the accessible part of a relation @{text R} given by
   259   Next we define the accessible part of a relation \<open>R\<close> given by
   263   the single rule:
   260   the single rule:
   264 
   261 
   265   \begin{center}\small
   262   \begin{center}\small
   266   \mbox{\inferrule{@{term "\<And>y. R y x \<Longrightarrow> accpart R y"}}{@{term "accpart R x"}}}
   263   \mbox{\inferrule{@{term "\<And>y. R y x \<Longrightarrow> accpart R y"}}{@{term "accpart R x"}}}
   267   \end{center}
   264   \end{center}
   268 
   265 
   269   The interesting point of this definition is that it contains a quantification
   266   The interesting point of this definition is that it contains a quantification
   270   that ranges only over the premise and the premise has also a precondition.
   267   that ranges only over the premise and the premise has also a precondition.
   271   The definition of @{text "accpart"} is:
   268   The definition of \<open>accpart\<close> is:
   272 *}
   269 \<close>
   273 
   270 
   274 definition "accpart \<equiv> \<lambda>R x. \<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P x"
   271 definition "accpart \<equiv> \<lambda>R x. \<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P x"
   275 
   272 
   276 text {*
   273 text \<open>
   277   The proof of the induction principle is again straightforward and omitted.
   274   The proof of the induction principle is again straightforward and omitted.
   278   Proving the introduction rule is a little more complicated, because the 
   275   Proving the introduction rule is a little more complicated, because the 
   279   quantifier and the implication in the premise. The proof is as follows.
   276   quantifier and the implication in the premise. The proof is as follows.
   280 *}
   277 \<close>
   281 
   278 
   282 lemma %linenos accpartI: 
   279 lemma %linenos accpartI: 
   283 shows "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
   280 shows "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
   284 apply (unfold accpart_def)
   281 apply (unfold accpart_def)
   285 apply (rule allI impI)+
   282 apply (rule allI impI)+
   297 	apply(rule p1[OF r1_prem, THEN spec[where x=P], THEN mp, OF r1])
   294 	apply(rule p1[OF r1_prem, THEN spec[where x=P], THEN mp, OF r1])
   298       done
   295       done
   299   qed
   296   qed
   300 qed
   297 qed
   301 
   298 
   302 text {*
   299 text \<open>
   303   As you can see, there are now two subproofs. The assumptions fall as usual into
   300   As you can see, there are now two subproofs. The assumptions fall as usual into
   304   two categories (Lines 7 to 9). In Line 11, applying the assumption @{text
   301   two categories (Lines 7 to 9). In Line 11, applying the assumption \<open>r1\<close> generates a goal state with the new local assumption @{term "R y x"},
   305   "r1"} generates a goal state with the new local assumption @{term "R y x"},
   302   named \<open>r1_prem\<close> in the second subproof (Line 14). This local assumption is
   306   named @{text "r1_prem"} in the second subproof (Line 14). This local assumption is
   303   used to solve the goal @{term "P y"} with the help of assumption \<open>p1\<close>.
   307   used to solve the goal @{term "P y"} with the help of assumption @{text
       
   308   "p1"}.
       
   309 
   304 
   310 
   305 
   311   \begin{exercise}
   306   \begin{exercise}
   312   Give the definition for the freshness predicate for lambda-terms. The rules
   307   Give the definition for the freshness predicate for lambda-terms. The rules
   313   for this predicate are:
   308   for this predicate are:
   326   The point of all these examples is to get a feeling what the automatic
   321   The point of all these examples is to get a feeling what the automatic
   327   proofs should do in order to solve all inductive definitions we throw at
   322   proofs should do in order to solve all inductive definitions we throw at
   328   them.  This is usually the first step in writing a package. We next explain
   323   them.  This is usually the first step in writing a package. We next explain
   329   the parsing and typing part of the package.
   324   the parsing and typing part of the package.
   330 
   325 
   331 *}
   326 \<close>
   332 (*<*)end(*>*)
   327 (*<*)end(*>*)