src/Doc/Tutorial/Inductive/Even.thy
author wenzelm
Sun, 06 Jan 2019 15:04:34 +0100
changeset 69605 a96320074298
parent 69597 ff784d5a5bfb
child 72991 d0a0b74f0ad7
permissions -rw-r--r--
isabelle update -u path_cartouches;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48895
4cd4ef1ef4a4 prefer ML_file over old uses;
wenzelm
parents: 43564
diff changeset
     1
(*<*)theory Even imports Main begin
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
     2
ML_file \<open>../../antiquote_setup.ML\<close> 
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 42637
diff changeset
     3
(*>*)
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
     4
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
     5
section\<open>The Set of Even Numbers\<close>
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
     6
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
     7
text \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
     8
\index{even numbers!defining inductively|(}%
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
     9
The set of even numbers can be inductively defined as the least set
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    10
containing 0 and closed under the operation $+2$.  Obviously,
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    11
\emph{even} can also be expressed using the divides relation (\<open>dvd\<close>). 
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    12
We shall prove below that the two formulations coincide.  On the way we
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    13
shall examine the primary means of reasoning about inductively defined
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    14
sets: rule induction.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    15
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    16
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    17
subsection\<open>Making an Inductive Definition\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    18
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    19
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    20
Using \commdx{inductive\protect\_set}, we declare the constant \<open>even\<close> to be
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    21
a set of natural numbers with the desired properties.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    22
\<close>
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
    23
25330
15bf0f47a87d added inductive
nipkow
parents: 23928
diff changeset
    24
inductive_set even :: "nat set" where
15bf0f47a87d added inductive
nipkow
parents: 23928
diff changeset
    25
zero[intro!]: "0 \<in> even" |
15bf0f47a87d added inductive
nipkow
parents: 23928
diff changeset
    26
step[intro!]: "n \<in> even \<Longrightarrow> (Suc (Suc n)) \<in> even"
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
    27
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    28
text \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    29
An inductive definition consists of introduction rules.  The first one
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    30
above states that 0 is even; the second states that if $n$ is even, then so
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    31
is~$n+2$.  Given this declaration, Isabelle generates a fixed point
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    32
definition for \<^term>\<open>even\<close> and proves theorems about it,
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    33
thus following the definitional approach (see {\S}\ref{sec:definitional}).
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    34
These theorems
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    35
include the introduction rules specified in the declaration, an elimination
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    36
rule for case analysis and an induction rule.  We can refer to these
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    37
theorems by automatically-generated names.  Here are two examples:
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    38
@{named_thms[display,indent=0] even.zero[no_vars] (even.zero) even.step[no_vars] (even.step)}
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
    39
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    40
The introduction rules can be given attributes.  Here
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    41
both rules are specified as \isa{intro!},%
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    42
\index{intro"!@\isa {intro"!} (attribute)}
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    43
directing the classical reasoner to 
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    44
apply them aggressively. Obviously, regarding 0 as even is safe.  The
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    45
\<open>step\<close> rule is also safe because $n+2$ is even if and only if $n$ is
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    46
even.  We prove this equivalence later.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    47
\<close>
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
    48
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    49
subsection\<open>Using Introduction Rules\<close>
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
    50
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    51
text \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    52
Our first lemma states that numbers of the form $2\times k$ are even.
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    53
Introduction rules are used to show that specific values belong to the
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    54
inductive set.  Such proofs typically involve 
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    55
induction, perhaps over some other inductive set.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    56
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    57
11705
ac8ca15c556c fixed numerals;
wenzelm
parents: 10883
diff changeset
    58
lemma two_times_even[intro!]: "2*k \<in> even"
12328
7c4ec77a8715 *** empty log message ***
nipkow
parents: 11705
diff changeset
    59
apply (induct_tac k)
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    60
 apply auto
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    61
done
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    62
(*<*)
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    63
lemma "2*k \<in> even"
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    64
apply (induct_tac k)
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    65
(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    66
txt \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    67
\noindent
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    68
The first step is induction on the natural number \<open>k\<close>, which leaves
10883
2b9f87bf9113 updated for new version of even-examples.tex
paulson
parents: 10341
diff changeset
    69
two subgoals:
2b9f87bf9113 updated for new version of even-examples.tex
paulson
parents: 10341
diff changeset
    70
@{subgoals[display,indent=0,margin=65]}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    71
Here \<open>auto\<close> simplifies both subgoals so that they match the introduction
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    72
rules, which are then applied automatically.
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
    73
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    74
Our ultimate goal is to prove the equivalence between the traditional
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    75
definition of \<open>even\<close> (using the divides relation) and our inductive
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    76
definition.  One direction of this equivalence is immediate by the lemma
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    77
just proved, whose \<open>intro!\<close> attribute ensures it is applied automatically.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    78
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    79
(*<*)oops(*>*)
11705
ac8ca15c556c fixed numerals;
wenzelm
parents: 10883
diff changeset
    80
lemma dvd_imp_even: "2 dvd n \<Longrightarrow> n \<in> even"
10883
2b9f87bf9113 updated for new version of even-examples.tex
paulson
parents: 10341
diff changeset
    81
by (auto simp add: dvd_def)
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
    82
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    83
subsection\<open>Rule Induction \label{sec:rule-induction}\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    84
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
    85
text \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    86
\index{rule induction|(}%
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    87
From the definition of the set
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    88
\<^term>\<open>even\<close>, Isabelle has
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    89
generated an induction rule:
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    90
@{named_thms [display,indent=0,margin=40] even.induct [no_vars] (even.induct)}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    91
A property \<^term>\<open>P\<close> holds for every even number provided it
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    92
holds for~\<open>0\<close> and is closed under the operation
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    93
\isa{Suc(Suc \(\cdot\))}.  Then \<^term>\<open>P\<close> is closed under the introduction
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    94
rules for \<^term>\<open>even\<close>, which is the least set closed under those rules. 
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    95
This type of inductive argument is called \textbf{rule induction}. 
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    96
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    97
Apart from the double application of \<^term>\<open>Suc\<close>, the induction rule above
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    98
resembles the familiar mathematical induction, which indeed is an instance
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
    99
of rule induction; the natural numbers can be defined inductively to be
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   100
the least set containing \<open>0\<close> and closed under~\<^term>\<open>Suc\<close>.
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   101
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   102
Induction is the usual way of proving a property of the elements of an
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   103
inductively defined set.  Let us prove that all members of the set
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   104
\<^term>\<open>even\<close> are multiples of two.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   105
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   106
11705
ac8ca15c556c fixed numerals;
wenzelm
parents: 10883
diff changeset
   107
lemma even_imp_dvd: "n \<in> even \<Longrightarrow> 2 dvd n"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   108
txt \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   109
We begin by applying induction.  Note that \<open>even.induct\<close> has the form
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   110
of an elimination rule, so we use the method \<open>erule\<close>.  We get two
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   111
subgoals:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   112
\<close>
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   113
apply (erule even.induct)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   114
txt \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   115
@{subgoals[display,indent=0]}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   116
We unfold the definition of \<open>dvd\<close> in both subgoals, proving the first
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   117
one and simplifying the second:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   118
\<close>
10883
2b9f87bf9113 updated for new version of even-examples.tex
paulson
parents: 10341
diff changeset
   119
apply (simp_all add: dvd_def)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   120
txt \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   121
@{subgoals[display,indent=0]}
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   122
The next command eliminates the existential quantifier from the assumption
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   123
and replaces \<open>n\<close> by \<open>2 * k\<close>.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   124
\<close>
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   125
apply clarify
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   126
txt \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   127
@{subgoals[display,indent=0]}
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   128
To conclude, we tell Isabelle that the desired value is
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   129
\<^term>\<open>Suc k\<close>.  With this hint, the subgoal falls to \<open>simp\<close>.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   130
\<close>
10883
2b9f87bf9113 updated for new version of even-examples.tex
paulson
parents: 10341
diff changeset
   131
apply (rule_tac x = "Suc k" in exI, simp)
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   132
(*<*)done(*>*)
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   133
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   134
text \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   135
Combining the previous two results yields our objective, the
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   136
equivalence relating \<^term>\<open>even\<close> and \<open>dvd\<close>. 
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   137
%
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   138
%we don't want [iff]: discuss?
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   139
\<close>
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   140
11705
ac8ca15c556c fixed numerals;
wenzelm
parents: 10883
diff changeset
   141
theorem even_iff_dvd: "(n \<in> even) = (2 dvd n)"
10883
2b9f87bf9113 updated for new version of even-examples.tex
paulson
parents: 10341
diff changeset
   142
by (blast intro: dvd_imp_even even_imp_dvd)
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   143
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   144
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   145
subsection\<open>Generalization and Rule Induction \label{sec:gen-rule-induction}\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   146
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   147
text \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   148
\index{generalizing for induction}%
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   149
Before applying induction, we typically must generalize
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   150
the induction formula.  With rule induction, the required generalization
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   151
can be hard to find and sometimes requires a complete reformulation of the
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   152
problem.  In this  example, our first attempt uses the obvious statement of
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   153
the result.  It fails:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   154
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   155
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   156
lemma "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even"
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   157
apply (erule even.induct)
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   158
oops
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   159
(*<*)
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   160
lemma "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even"
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   161
apply (erule even.induct)
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   162
(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   163
txt \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   164
Rule induction finds no occurrences of \<^term>\<open>Suc(Suc n)\<close> in the
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   165
conclusion, which it therefore leaves unchanged.  (Look at
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   166
\<open>even.induct\<close> to see why this happens.)  We have these subgoals:
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   167
@{subgoals[display,indent=0]}
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   168
The first one is hopeless.  Rule induction on
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   169
a non-variable term discards information, and usually fails.
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   170
How to deal with such situations
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   171
in general is described in {\S}\ref{sec:ind-var-in-prems} below.
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   172
In the current case the solution is easy because
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   173
we have the necessary inverse, subtraction:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   174
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   175
(*<*)oops(*>*)
11705
ac8ca15c556c fixed numerals;
wenzelm
parents: 10883
diff changeset
   176
lemma even_imp_even_minus_2: "n \<in> even \<Longrightarrow> n - 2 \<in> even"
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   177
apply (erule even.induct)
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   178
 apply auto
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   179
done
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   180
(*<*)
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   181
lemma "n \<in>  even \<Longrightarrow> n - 2 \<in> even"
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   182
apply (erule even.induct)
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   183
(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   184
txt \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   185
This lemma is trivially inductive.  Here are the subgoals:
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   186
@{subgoals[display,indent=0]}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   187
The first is trivial because \<open>0 - 2\<close> simplifies to \<open>0\<close>, which is
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   188
even.  The second is trivial too: \<^term>\<open>Suc (Suc n) - 2\<close> simplifies to
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   189
\<^term>\<open>n\<close>, matching the assumption.%
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   190
\index{rule induction|)}  %the sequel isn't really about induction
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   191
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   192
\medskip
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   193
Using our lemma, we can easily prove the result we originally wanted:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   194
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   195
(*<*)oops(*>*)
10883
2b9f87bf9113 updated for new version of even-examples.tex
paulson
parents: 10341
diff changeset
   196
lemma Suc_Suc_even_imp_even: "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even"
2b9f87bf9113 updated for new version of even-examples.tex
paulson
parents: 10341
diff changeset
   197
by (drule even_imp_even_minus_2, simp)
10326
d4fe5ce8a5d5 minor tinkering
paulson
parents: 10314
diff changeset
   198
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   199
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   200
We have just proved the converse of the introduction rule \<open>even.step\<close>.
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   201
This suggests proving the following equivalence.  We give it the
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   202
\attrdx{iff} attribute because of its obvious value for simplification.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   203
\<close>
10326
d4fe5ce8a5d5 minor tinkering
paulson
parents: 10314
diff changeset
   204
d4fe5ce8a5d5 minor tinkering
paulson
parents: 10314
diff changeset
   205
lemma [iff]: "((Suc (Suc n)) \<in> even) = (n \<in> even)"
10883
2b9f87bf9113 updated for new version of even-examples.tex
paulson
parents: 10341
diff changeset
   206
by (blast dest: Suc_Suc_even_imp_even)
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   207
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   208
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   209
subsection\<open>Rule Inversion \label{sec:rule-inversion}\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   210
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   211
text \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   212
\index{rule inversion|(}%
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   213
Case analysis on an inductive definition is called \textbf{rule
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   214
inversion}.  It is frequently used in proofs about operational
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   215
semantics.  It can be highly effective when it is applied
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   216
automatically.  Let us look at how rule inversion is done in
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   217
Isabelle/HOL\@.
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   218
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   219
Recall that \<^term>\<open>even\<close> is the minimal set closed under these two rules:
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   220
@{thm [display,indent=0] even.intros [no_vars]}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   221
Minimality means that \<^term>\<open>even\<close> contains only the elements that these
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   222
rules force it to contain.  If we are told that \<^term>\<open>a\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   223
belongs to
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   224
\<^term>\<open>even\<close> then there are only two possibilities.  Either \<^term>\<open>a\<close> is \<open>0\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   225
or else \<^term>\<open>a\<close> has the form \<^term>\<open>Suc(Suc n)\<close>, for some suitable \<^term>\<open>n\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   226
that belongs to
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   227
\<^term>\<open>even\<close>.  That is the gist of the \<^term>\<open>cases\<close> rule, which Isabelle proves
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   228
for us when it accepts an inductive definition:
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   229
@{named_thms [display,indent=0,margin=40] even.cases [no_vars] (even.cases)}
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   230
This general rule is less useful than instances of it for
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   231
specific patterns.  For example, if \<^term>\<open>a\<close> has the form
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   232
\<^term>\<open>Suc(Suc n)\<close> then the first case becomes irrelevant, while the second
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   233
case tells us that \<^term>\<open>n\<close> belongs to \<^term>\<open>even\<close>.  Isabelle will generate
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   234
this instance for us:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   235
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   236
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   237
inductive_cases Suc_Suc_cases [elim!]: "Suc(Suc n) \<in> even"
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   238
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   239
text \<open>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   240
The \commdx{inductive\protect\_cases} command generates an instance of
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   241
the \<open>cases\<close> rule for the supplied pattern and gives it the supplied name:
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   242
@{named_thms [display,indent=0] Suc_Suc_cases [no_vars] (Suc_Suc_cases)}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   243
Applying this as an elimination rule yields one case where \<open>even.cases\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   244
would yield two.  Rule inversion works well when the conclusions of the
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   245
introduction rules involve datatype constructors like \<^term>\<open>Suc\<close> and \<open>#\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   246
(list ``cons''); freeness reasoning discards all but one or two cases.
10314
5b36035e4dff even numbers example
paulson
parents:
diff changeset
   247
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   248
In the \isacommand{inductive\_cases} command we supplied an
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   249
attribute, \<open>elim!\<close>,
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   250
\index{elim"!@\isa {elim"!} (attribute)}%
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   251
indicating that this elimination rule can be
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   252
applied aggressively.  The original
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   253
\<^term>\<open>cases\<close> rule would loop if used in that manner because the
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   254
pattern~\<^term>\<open>a\<close> matches everything.
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   255
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   256
The rule \<open>Suc_Suc_cases\<close> is equivalent to the following implication:
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   257
@{term [display,indent=0] "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even"}
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   258
Just above we devoted some effort to reaching precisely
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   259
this result.  Yet we could have obtained it by a one-line declaration,
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   260
dispensing with the lemma \<open>even_imp_even_minus_2\<close>. 
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   261
This example also justifies the terminology
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   262
\textbf{rule inversion}: the new rule inverts the introduction rule
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   263
\<open>even.step\<close>.  In general, a rule can be inverted when the set of elements
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   264
it introduces is disjoint from those of the other introduction rules.
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   265
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   266
For one-off applications of rule inversion, use the \methdx{ind_cases} method. 
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   267
Here is an example:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   268
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   269
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   270
(*<*)lemma "Suc(Suc n) \<in> even \<Longrightarrow> P"(*>*)
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   271
apply (ind_cases "Suc(Suc n) \<in> even")
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   272
(*<*)oops(*>*)
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   273
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   274
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   275
The specified instance of the \<open>cases\<close> rule is generated, then applied
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   276
as an elimination rule.
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   277
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   278
To summarize, every inductive definition produces a \<open>cases\<close> rule.  The
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   279
\commdx{inductive\protect\_cases} command stores an instance of the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   280
\<open>cases\<close> rule for a given pattern.  Within a proof, the
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   281
\<open>ind_cases\<close> method applies an instance of the \<open>cases\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   282
rule.
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   283
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   284
The even numbers example has shown how inductive definitions can be
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   285
used.  Later examples will show that they are actually worth using.%
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   286
\index{rule inversion|)}%
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   287
\index{even numbers!defining inductively|)}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 56059
diff changeset
   288
\<close>
23842
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   289
9d87177f1f89 LaTeX code is now generated directly from theory file.
berghofe
parents: 23733
diff changeset
   290
(*<*)end(*>*)